doc-src/TutorialI/Advanced/WFrec.thy
author chaieb
Wed, 19 May 2004 11:23:59 +0200
changeset 14758 af3b71a46a1c
parent 13111 2d6782e71702
child 17914 99ead7a7eb42
permissions -rw-r--r--
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller. the tactic has also changed and allows the abstaction over fuction occurences whose type is nat or int.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10187
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
     1
(*<*)theory WFrec = Main:(*>*)
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
     2
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
     3
text{*\noindent
11161
166f7d87b37f *** empty log message ***
nipkow
parents: 10885
diff changeset
     4
So far, all recursive definitions were shown to terminate via measure
11494
23a118849801 revisions and indexing
paulson
parents: 11429
diff changeset
     5
functions. Sometimes this can be inconvenient or
10187
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
     6
impossible. Fortunately, \isacommand{recdef} supports much more
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
     7
general definitions. For example, termination of Ackermann's function
10654
458068404143 *** empty log message ***
nipkow
parents: 10545
diff changeset
     8
can be shown by means of the \rmindex{lexicographic product} @{text"<*lex*>"}:
10187
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
     9
*}
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    10
11626
0dbfb578bf75 recdef (permissive);
wenzelm
parents: 11494
diff changeset
    11
consts ack :: "nat\<times>nat \<Rightarrow> nat"
10187
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    12
recdef ack "measure(\<lambda>m. m) <*lex*> measure(\<lambda>n. n)"
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    13
  "ack(0,n)         = Suc n"
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    14
  "ack(Suc m,0)     = ack(m, 1)"
11626
0dbfb578bf75 recdef (permissive);
wenzelm
parents: 11494
diff changeset
    15
  "ack(Suc m,Suc n) = ack(m,ack(Suc m,n))"
10187
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    16
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    17
text{*\noindent
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    18
The lexicographic product decreases if either its first component
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    19
decreases (as in the second equation and in the outer call in the
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    20
third equation) or its first component stays the same and the second
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    21
component decreases (as in the inner call in the third equation).
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    22
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    23
In general, \isacommand{recdef} supports termination proofs based on
10396
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
    24
arbitrary well-founded relations as introduced in \S\ref{sec:Well-founded}.
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
    25
This is called \textbf{well-founded
11494
23a118849801 revisions and indexing
paulson
parents: 11429
diff changeset
    26
recursion}\indexbold{recursion!well-founded}.  A function definition
23a118849801 revisions and indexing
paulson
parents: 11429
diff changeset
    27
is total if and only if the set of 
23a118849801 revisions and indexing
paulson
parents: 11429
diff changeset
    28
all pairs $(r,l)$, where $l$ is the argument on the
10396
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
    29
left-hand side of an equation and $r$ the argument of some recursive call on
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
    30
the corresponding right-hand side, induces a well-founded relation.  For a
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
    31
systematic account of termination proofs via well-founded relations see, for
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10841
diff changeset
    32
example, Baader and Nipkow~\cite{Baader-Nipkow}.
10187
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    33
11494
23a118849801 revisions and indexing
paulson
parents: 11429
diff changeset
    34
Each \isacommand{recdef} definition should be accompanied (after the function's
23a118849801 revisions and indexing
paulson
parents: 11429
diff changeset
    35
name) by a well-founded relation on the function's argument type.  
23a118849801 revisions and indexing
paulson
parents: 11429
diff changeset
    36
Isabelle/HOL formalizes some of the most important
10396
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
    37
constructions of well-founded relations (see \S\ref{sec:Well-founded}). For
11494
23a118849801 revisions and indexing
paulson
parents: 11429
diff changeset
    38
example, @{term"measure f"} is always well-founded.   The lexicographic
10396
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
    39
product of two well-founded relations is again well-founded, which we relied
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
    40
on when defining Ackermann's function above.
11308
b28bbb153603 *** empty log message ***
nipkow
parents: 11196
diff changeset
    41
Of course the lexicographic product can also be iterated:
10189
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    42
*}
10187
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    43
10189
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    44
consts contrived :: "nat \<times> nat \<times> nat \<Rightarrow> nat"
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    45
recdef contrived
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    46
  "measure(\<lambda>i. i) <*lex*> measure(\<lambda>j. j) <*lex*> measure(\<lambda>k. k)"
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    47
"contrived(i,j,Suc k) = contrived(i,j,k)"
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    48
"contrived(i,Suc j,0) = contrived(i,j,j)"
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    49
"contrived(Suc i,0,0) = contrived(i,i,i)"
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    50
"contrived(0,0,0)     = 0"
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    51
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    52
text{*
10396
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
    53
Lexicographic products of measure functions already go a long
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10841
diff changeset
    54
way. Furthermore, you may embed a type in an
10396
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
    55
existing well-founded relation via the inverse image construction @{term
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
    56
inv_image}. All these constructions are known to \isacommand{recdef}. Thus you
10241
e0428c2778f1 wellfounded -> well-founded
paulson
parents: 10190
diff changeset
    57
will never have to prove well-foundedness of any relation composed
10189
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    58
solely of these building blocks. But of course the proof of
11494
23a118849801 revisions and indexing
paulson
parents: 11429
diff changeset
    59
termination of your function definition --- that the arguments
23a118849801 revisions and indexing
paulson
parents: 11429
diff changeset
    60
decrease with every recursive call --- may still require you to provide
10189
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    61
additional lemmas.
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    62
10841
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    63
It is also possible to use your own well-founded relations with
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    64
\isacommand{recdef}.  For example, the greater-than relation can be made
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    65
well-founded by cutting it off at a certain point.  Here is an example
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    66
of a recursive function that calls itself with increasing values up to ten:
10187
0376cccd9118 *** empty log message ***
nipkow
parents:
diff changeset
    67
*}
10189
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    68
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    69
consts f :: "nat \<Rightarrow> nat"
11705
ac8ca15c556c fixed numerals;
wenzelm
parents: 11636
diff changeset
    70
recdef (*<*)(permissive)(*>*)f "{(i,j). j<i \<and> i \<le> (10::nat)}"
ac8ca15c556c fixed numerals;
wenzelm
parents: 11636
diff changeset
    71
"f i = (if 10 \<le> i then 0 else i * f(Suc i))"
10189
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    72
10396
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
    73
text{*\noindent
10841
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    74
Since \isacommand{recdef} is not prepared for the relation supplied above,
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    75
Isabelle rejects the definition.  We should first have proved that
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    76
our relation was well-founded:
10189
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    77
*}
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
    78
10841
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    79
lemma wf_greater: "wf {(i,j). j<i \<and> i \<le> (N::nat)}"
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    80
11196
bb4ede27fcb7 *** empty log message ***
nipkow
parents: 11161
diff changeset
    81
txt{*\noindent
10841
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    82
The proof is by showing that our relation is a subset of another well-founded
11494
23a118849801 revisions and indexing
paulson
parents: 11429
diff changeset
    83
relation: one given by a measure function.\index{*wf_subset (theorem)}
11626
0dbfb578bf75 recdef (permissive);
wenzelm
parents: 11494
diff changeset
    84
*}
10841
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    85
11626
0dbfb578bf75 recdef (permissive);
wenzelm
parents: 11494
diff changeset
    86
apply (rule wf_subset [of "measure (\<lambda>k::nat. N-k)"], blast)
10841
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    87
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    88
txt{*
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    89
@{subgoals[display,indent=0,margin=65]}
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    90
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    91
\noindent
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    92
The inclusion remains to be proved. After unfolding some definitions, 
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    93
we are left with simple arithmetic:
11626
0dbfb578bf75 recdef (permissive);
wenzelm
parents: 11494
diff changeset
    94
*}
10841
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    95
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    96
apply (clarify, simp add: measure_def inv_image_def)
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    97
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    98
txt{*
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
    99
@{subgoals[display,indent=0,margin=65]}
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   100
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   101
\noindent
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   102
And that is dispatched automatically:
11626
0dbfb578bf75 recdef (permissive);
wenzelm
parents: 11494
diff changeset
   103
*}
10841
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   104
11626
0dbfb578bf75 recdef (permissive);
wenzelm
parents: 11494
diff changeset
   105
by arith
10189
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
   106
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
   107
text{*\noindent
10841
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   108
11429
30da2f5eaf57 tidying the index
paulson
parents: 11308
diff changeset
   109
Armed with this lemma, we use the \attrdx{recdef_wf} attribute to attach a
13111
2d6782e71702 *** empty log message ***
nipkow
parents: 11705
diff changeset
   110
crucial hint\cmmdx{hints} to our definition:
10189
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
   111
*}
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
   112
(*<*)
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
   113
consts g :: "nat \<Rightarrow> nat"
11705
ac8ca15c556c fixed numerals;
wenzelm
parents: 11636
diff changeset
   114
recdef g "{(i,j). j<i \<and> i \<le> (10::nat)}"
ac8ca15c556c fixed numerals;
wenzelm
parents: 11636
diff changeset
   115
"g i = (if 10 \<le> i then 0 else i * g(Suc i))"
10189
865918597b63 *** empty log message ***
nipkow
parents: 10187
diff changeset
   116
(*>*)
11626
0dbfb578bf75 recdef (permissive);
wenzelm
parents: 11494
diff changeset
   117
(hints recdef_wf: wf_greater)
10841
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   118
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   119
text{*\noindent
11705
ac8ca15c556c fixed numerals;
wenzelm
parents: 11636
diff changeset
   120
Alternatively, we could have given @{text "measure (\<lambda>k::nat. 10-k)"} for the
10841
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   121
well-founded relation in our \isacommand{recdef}.  However, the arithmetic
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   122
goal in the lemma above would have arisen instead in the \isacommand{recdef}
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   123
termination proof, where we have less control.  A tailor-made termination
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   124
relation makes even more sense when it can be used in several function
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   125
declarations.
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   126
*}
2fb8089ab6cd new wfrec example
paulson
parents: 10795
diff changeset
   127
10396
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10241
diff changeset
   128
(*<*)end(*>*)