doc-src/TutorialI/Recdef/simplification.thy
author wenzelm
Tue, 09 May 2000 15:10:25 +0200
changeset 8845 03a2ae3059da
parent 8771 026f37a86ea7
child 9458 c613cd06d5cf
permissions -rw-r--r--
updated keywords;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     1
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     2
theory simplification = Main:;
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{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     6
Once we have succeeded in proving all termination conditions, the recursion
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     7
equations become simplification rules, just as with
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     8
\isacommand{primrec}. In most cases this works fine, but there is a subtle
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     9
problem that must be mentioned: simplification may not
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    10
terminate because of automatic splitting of \isa{if}.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    11
Let us look at an example:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    12
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    13
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    14
consts gcd :: "nat*nat \\<Rightarrow> nat";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
recdef gcd "measure (\\<lambda>(m,n).n)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    16
  "gcd (m, n) = (if n=0 then m else gcd(n, m mod n))";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    17
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
According to the measure function, the second argument should decrease with
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
each recursive call. The resulting termination condition
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    21
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    23
(*<*)term(*>*) "n \\<noteq> 0 \\<Longrightarrow> m mod n < n";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    24
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    25
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    26
is provded automatically because it is already present as a lemma in the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    27
arithmetic library. Thus the recursion equation becomes a simplification
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    28
rule. Of course the equation is nonterminating if we are allowed to unfold
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    29
the recursive call inside the \isa{else} branch, which is why programming
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    30
languages and our simplifier don't do that. Unfortunately the simplifier does
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    31
something else which leads to the same problem: it splits \isa{if}s if the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    32
condition simplifies to neither \isa{True} nor \isa{False}. For
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
example, simplification reduces
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    34
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    35
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    36
(*<*)term(*>*) "gcd(m,n) = k";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    37
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    38
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    39
in one step to
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    40
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    42
(*<*)term(*>*) "(if n=0 then m else gcd(n, m mod n)) = k";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    43
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    45
where the condition cannot be reduced further, and splitting leads to
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    46
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    47
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
(*<*)term(*>*) "(n=0 \\<longrightarrow> m=k) \\<and> (n\\<noteq>0 \\<longrightarrow> gcd(n, m mod n)=k)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    49
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    50
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    51
Since the recursive call \isa{gcd(n, m mod n)} is no longer protected by
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    52
an \isa{if}, it is unfolded again, which leads to an infinite chain of simplification steps.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    53
Fortunately, this problem can be avoided in many different ways.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    54
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    55
The most radical solution is to disable the offending
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    56
\isa{split_if} as shown in the section on case splits in
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    57
\S\ref{sec:SimpFeatures}.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    58
However, we do not recommend this because it means you will often have to
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    59
invoke the rule explicitly when \isa{if} is involved.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    60
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    61
If possible, the definition should be given by pattern matching on the left
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    62
rather than \isa{if} on the right. In the case of \isa{gcd} the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    63
following alternative definition suggests itself:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    64
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    65
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    66
consts gcd1 :: "nat*nat \\<Rightarrow> nat";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    67
recdef gcd1 "measure (\\<lambda>(m,n).n)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    68
  "gcd1 (m, 0) = m"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    69
  "gcd1 (m, n) = gcd1(n, m mod n)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    70
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    71
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    72
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    73
Note that the order of equations is important and hides the side condition
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    74
\isa{n \isasymnoteq\ 0}. Unfortunately, in general the case distinction
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    75
may not be expressible by pattern matching.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    76
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    77
A very simple alternative is to replace \isa{if} by \isa{case}, which
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    78
is also available for \isa{bool} but is not split automatically:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    79
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    80
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    81
consts gcd2 :: "nat*nat \\<Rightarrow> nat";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    82
recdef gcd2 "measure (\\<lambda>(m,n).n)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    83
  "gcd2(m,n) = (case n=0 of True \\<Rightarrow> m | False \\<Rightarrow> gcd2(n,m mod n))";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    84
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    85
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    86
In fact, this is probably the neatest solution next to pattern matching.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    87
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    88
A final alternative is to replace the offending simplification rules by
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    89
derived conditional ones. For \isa{gcd} it means we have to prove
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    90
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    91
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    92
lemma [simp]: "gcd (m, 0) = m";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    93
apply(simp).;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    94
lemma [simp]: "n \\<noteq> 0 \\<Longrightarrow> gcd(m, n) = gcd(n, m mod n)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    95
apply(simp).;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    96
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    97
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    98
after which we can disable the original simplification rule:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    99
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   100
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   101
lemmas [simp del] = gcd.simps;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   102
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   103
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   104
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   105
(*>*)