doc-src/TutorialI/Recdef/termination.thy
author paulson
Fri, 03 Aug 2001 18:04:55 +0200
changeset 11458 09a6c44a48ea
parent 11429 30da2f5eaf57
child 11626 0dbfb578bf75
permissions -rw-r--r--
numerous stylistic changes and indexing
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     1
(*<*)
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9458
diff changeset
     2
theory termination = examples:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     3
(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     4
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     5
text{*
11309
d666f11ca2d4 minor suggestions by Tanja Vos
paulson
parents: 10971
diff changeset
     6
When a function~$f$ is defined via \isacommand{recdef}, Isabelle tries to prove
d666f11ca2d4 minor suggestions by Tanja Vos
paulson
parents: 10971
diff changeset
     7
its termination with the help of the user-supplied measure.  Each of the examples
d666f11ca2d4 minor suggestions by Tanja Vos
paulson
parents: 10971
diff changeset
     8
above is simple enough that Isabelle can automatically prove that the
d666f11ca2d4 minor suggestions by Tanja Vos
paulson
parents: 10971
diff changeset
     9
argument's measure decreases in each recursive call. As a result,
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9458
diff changeset
    10
$f$@{text".simps"} will contain the defining equations (or variants derived
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9458
diff changeset
    11
from them) as theorems. For example, look (via \isacommand{thm}) at
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9458
diff changeset
    12
@{thm[source]sep.simps} and @{thm[source]sep1.simps} to see that they define
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9458
diff changeset
    13
the same function. What is more, those equations are automatically declared as
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    14
simplification rules.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11429
diff changeset
    16
Isabelle may fail to prove the termination condition for some
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11429
diff changeset
    17
recursive call.  Let us try the following artificial function:*}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    19
consts f :: "nat\<times>nat \<Rightarrow> nat";
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    20
recdef f "measure(\<lambda>(x,y). x-y)"
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    21
  "f(x,y) = (if x \<le> y then x else f(x,y+1))";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    23
text{*\noindent
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11429
diff changeset
    24
Isabelle prints a
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11429
diff changeset
    25
\REMARK{error or warning?  change this part?  rename g to f?}
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10654
diff changeset
    26
message showing you what it was unable to prove. You will then
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    27
have to prove it as a separate lemma before you attempt the definition
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    28
of your function once more. In our case the required lemma is the obvious one:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    29
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    30
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    31
lemma termi_lem: "\<not> x \<le> y \<Longrightarrow> x - Suc y < x - y";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    32
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
txt{*\noindent
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11429
diff changeset
    34
It was not proved automatically because of the awkward behaviour of subtraction
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11429
diff changeset
    35
on type @{typ"nat"}. This requires more arithmetic than is tried by default:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    36
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    37
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9992
diff changeset
    38
apply(arith);
59d6633835fa *** empty log message ***
nipkow
parents: 9992
diff changeset
    39
done
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    40
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
text{*\noindent
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    42
Because \isacommand{recdef}'s termination prover involves simplification,
11429
30da2f5eaf57 tidying the index
paulson
parents: 11309
diff changeset
    43
we include in our second attempt a hint: the \attrdx{recdef_simp} attribute 
30da2f5eaf57 tidying the index
paulson
parents: 11309
diff changeset
    44
says to use @{thm[source]termi_lem} as
30da2f5eaf57 tidying the index
paulson
parents: 11309
diff changeset
    45
a simplification rule.  
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    46
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    47
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    48
consts g :: "nat\<times>nat \<Rightarrow> nat";
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    49
recdef g "measure(\<lambda>(x,y). x-y)"
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    50
  "g(x,y) = (if x \<le> y then x else g(x,y+1))"
9992
4281ccea43f0 *** empty log message ***
nipkow
parents: 9933
diff changeset
    51
(hints recdef_simp: termi_lem)
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    52
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    53
text{*\noindent
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9458
diff changeset
    54
This time everything works fine. Now @{thm[source]g.simps} contains precisely
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11429
diff changeset
    55
the stated recursion equation for @{term g}, which has been stored as a
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11429
diff changeset
    56
simplification rule.  Thus we can automatically prove results such as this one:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    57
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    58
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    59
theorem "g(1,0) = g(1,1)";
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9992
diff changeset
    60
apply(simp);
59d6633835fa *** empty log message ***
nipkow
parents: 9992
diff changeset
    61
done
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    62
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    63
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    64
More exciting theorems require induction, which is discussed below.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    65
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    66
If the termination proof requires a new lemma that is of general use, you can
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    67
turn it permanently into a simplification rule, in which case the above
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    68
\isacommand{hint} is not necessary. But our @{thm[source]termi_lem} is not
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    69
sufficiently general to warrant this distinction.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    70
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9992
diff changeset
    71
The attentive reader may wonder why we chose to call our function @{term g}
59d6633835fa *** empty log message ***
nipkow
parents: 9992
diff changeset
    72
rather than @{term f} the second time around. The reason is that, despite
59d6633835fa *** empty log message ***
nipkow
parents: 9992
diff changeset
    73
the failed termination proof, the definition of @{term f} did not
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9458
diff changeset
    74
fail, and thus we could not define it a second time. However, all theorems
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9992
diff changeset
    75
about @{term f}, for example @{thm[source]f.simps}, carry as a precondition
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9458
diff changeset
    76
the unproved termination condition. Moreover, the theorems
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11429
diff changeset
    77
@{thm[source]f.simps} are not stored as simplification rules. 
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11429
diff changeset
    78
However, this mechanism
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9458
diff changeset
    79
allows a delayed proof of termination: instead of proving
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9458
diff changeset
    80
@{thm[source]termi_lem} up front, we could prove 
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    81
it later on and then use it to remove the preconditions from the theorems
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9992
diff changeset
    82
about @{term f}. In most cases this is more cumbersome than proving things
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9458
diff changeset
    83
up front.
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11429
diff changeset
    84
\REMARK{FIXME, with one exception: nested recursion.}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    85
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    86
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    87
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    88
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    89
(*>*)