src/Doc/Tutorial/Recdef/simplification.thy
author wenzelm
Thu, 28 Dec 2017 12:26:57 +0100
changeset 67286 417e081322ae
parent 62392 747d36865c2c
child 67406 23307fd33906
permissions -rw-r--r--
avoid clash with special files in HTML output;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     1
(*<*)
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 48985
diff changeset
     2
theory simplification imports Main begin
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{*
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
     6
Once we have proved all the termination conditions, the \isacommand{recdef} 
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
     7
recursion equations become simplification rules, just as with
8745
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
19792
e8e3da6d3ff7 quoted "if";
wenzelm
parents: 16417
diff changeset
    10
terminate because of automatic splitting of @{text "if"}.
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
    11
\index{*if expressions!splitting of}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    12
Let us look at an example:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    13
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    14
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 48985
diff changeset
    15
consts gcd :: "nat\<times>nat \<Rightarrow> nat"
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9933
diff changeset
    16
recdef gcd "measure (\<lambda>(m,n).n)"
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 48985
diff changeset
    17
  "gcd (m, n) = (if n=0 then m else gcd(n, m mod n))"
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
According to the measure function, the second argument should decrease with
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    21
each recursive call. The resulting termination condition
12473
f41e477576b9 *** empty log message ***
nipkow
parents: 11458
diff changeset
    22
@{term[display]"n ~= (0::nat) ==> m mod n < n"}
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10795
diff changeset
    23
is proved automatically because it is already present as a lemma in
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10795
diff changeset
    24
HOL\@.  Thus the recursion equation becomes a simplification
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    25
rule. Of course the equation is nonterminating if we are allowed to unfold
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9933
diff changeset
    26
the recursive call inside the @{text else} branch, which is why programming
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    27
languages and our simplifier don't do that. Unfortunately the simplifier does
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
    28
something else that leads to the same problem: it splits 
19792
e8e3da6d3ff7 quoted "if";
wenzelm
parents: 16417
diff changeset
    29
each @{text "if"}-expression unless its
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
    30
condition simplifies to @{term True} or @{term False}.  For
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    31
example, simplification reduces
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9458
diff changeset
    32
@{term[display]"gcd(m,n) = k"}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
in one step to
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9458
diff changeset
    34
@{term[display]"(if n=0 then m else gcd(n, m mod n)) = k"}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    35
where the condition cannot be reduced further, and splitting leads to
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9458
diff changeset
    36
@{term[display]"(n=0 --> m=k) & (n ~= 0 --> gcd(n, m mod n)=k)"}
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9458
diff changeset
    37
Since the recursive call @{term"gcd(n, m mod n)"} is no longer protected by
19792
e8e3da6d3ff7 quoted "if";
wenzelm
parents: 16417
diff changeset
    38
an @{text "if"}, it is unfolded again, which leads to an infinite chain of
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9458
diff changeset
    39
simplification steps. Fortunately, this problem can be avoided in many
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9458
diff changeset
    40
different ways.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
    42
The most radical solution is to disable the offending theorem
62392
747d36865c2c more canonical names
nipkow
parents: 58860
diff changeset
    43
@{thm[source]if_split},
11215
b44ad7e4c4d2 *** empty log message ***
nipkow
parents: 10885
diff changeset
    44
as shown in \S\ref{sec:AutoCaseSplits}.  However, we do not recommend this
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
    45
approach: you will often have to invoke the rule explicitly when
19792
e8e3da6d3ff7 quoted "if";
wenzelm
parents: 16417
diff changeset
    46
@{text "if"} is involved.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    47
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
If possible, the definition should be given by pattern matching on the left
19792
e8e3da6d3ff7 quoted "if";
wenzelm
parents: 16417
diff changeset
    49
rather than @{text "if"} on the right. In the case of @{term gcd} the
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    50
following alternative definition suggests itself:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    51
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    52
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 48985
diff changeset
    53
consts gcd1 :: "nat\<times>nat \<Rightarrow> nat"
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9933
diff changeset
    54
recdef gcd1 "measure (\<lambda>(m,n).n)"
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    55
  "gcd1 (m, 0) = m"
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 48985
diff changeset
    56
  "gcd1 (m, n) = gcd1(n, m mod n)"
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
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    59
text{*\noindent
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
    60
The order of equations is important: it hides the side condition
12473
f41e477576b9 *** empty log message ***
nipkow
parents: 11458
diff changeset
    61
@{prop"n ~= (0::nat)"}.  Unfortunately, in general the case distinction
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    62
may not be expressible by pattern matching.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    63
19792
e8e3da6d3ff7 quoted "if";
wenzelm
parents: 16417
diff changeset
    64
A simple alternative is to replace @{text "if"} by @{text case}, 
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
    65
which is also available for @{typ bool} and is not split automatically:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    66
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    67
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 48985
diff changeset
    68
consts gcd2 :: "nat\<times>nat \<Rightarrow> nat"
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9933
diff changeset
    69
recdef gcd2 "measure (\<lambda>(m,n).n)"
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 48985
diff changeset
    70
  "gcd2(m,n) = (case n=0 of True \<Rightarrow> m | False \<Rightarrow> gcd2(n,m mod n))"
8745
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
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
    73
This is probably the neatest solution next to pattern matching, and it is
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
    74
always available.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    75
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    76
A final alternative is to replace the offending simplification rules by
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9933
diff changeset
    77
derived conditional ones. For @{term gcd} it means we have to prove
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10178
diff changeset
    78
these lemmas:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    79
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    80
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 48985
diff changeset
    81
lemma [simp]: "gcd (m, 0) = m"
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 48985
diff changeset
    82
apply(simp)
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9933
diff changeset
    83
done
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
    84
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 48985
diff changeset
    85
lemma [simp]: "n \<noteq> 0 \<Longrightarrow> gcd(m, n) = gcd(n, m mod n)"
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 48985
diff changeset
    86
apply(simp)
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9933
diff changeset
    87
done
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    88
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    89
text{*\noindent
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11215
diff changeset
    90
Simplification terminates for these proofs because the condition of the @{text
19792
e8e3da6d3ff7 quoted "if";
wenzelm
parents: 16417
diff changeset
    91
"if"} simplifies to @{term True} or @{term False}.
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10178
diff changeset
    92
Now we can disable the original simplification rule:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    93
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    94
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9834
diff changeset
    95
declare gcd.simps [simp del]
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    96
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    97
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    98
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    99
(*>*)