doc-src/TutorialI/Recdef/simplification.thy
changeset 9792 bbefb6ce5cb2
parent 9754 a123a64cadeb
child 9834 109b11c4e77e
--- a/doc-src/TutorialI/Recdef/simplification.thy	Fri Sep 01 18:29:52 2000 +0200
+++ b/doc-src/TutorialI/Recdef/simplification.thy	Fri Sep 01 19:09:44 2000 +0200
@@ -7,7 +7,7 @@
 equations become simplification rules, just as with
 \isacommand{primrec}. In most cases this works fine, but there is a subtle
 problem that must be mentioned: simplification may not
-terminate because of automatic splitting of @{name"if"}.
+terminate because of automatic splitting of @{text"if"}.
 Let us look at an example:
 *}
 
@@ -24,9 +24,9 @@
 is provded automatically because it is already present as a lemma in the
 arithmetic library. Thus the recursion equation becomes a simplification
 rule. Of course the equation is nonterminating if we are allowed to unfold
-the recursive call inside the @{name"if"} branch, which is why programming
+the recursive call inside the @{text"if"} branch, which is why programming
 languages and our simplifier don't do that. Unfortunately the simplifier does
-something else which leads to the same problem: it splits @{name"if"}s if the
+something else which leads to the same problem: it splits @{text"if"}s if the
 condition simplifies to neither @{term"True"} nor @{term"False"}. For
 example, simplification reduces
 \begin{quote}
@@ -41,17 +41,18 @@
 @{term[display]"(n=0 --> m=k) & (n ~= 0 --> gcd(n, m mod n)=k)"}
 \end{quote}
 Since the recursive call @{term"gcd(n, m mod n)"} is no longer protected by
-an @{name"if"}, it is unfolded again, which leads to an infinite chain of
+an @{text"if"}, it is unfolded again, which leads to an infinite chain of
 simplification steps. Fortunately, this problem can be avoided in many
 different ways.
 
-The most radical solution is to disable the offending \@{name"split_if"} as
-shown in the section on case splits in \S\ref{sec:Simplification}.  However,
-we do not recommend this because it means you will often have to invoke the
-rule explicitly when @{name"if"} is involved.
+The most radical solution is to disable the offending
+@{thm[source]split_if} as shown in the section on case splits in
+\S\ref{sec:Simplification}.  However, we do not recommend this because it
+means you will often have to invoke the rule explicitly when @{text"if"} is
+involved.
 
 If possible, the definition should be given by pattern matching on the left
-rather than @{name"if"} on the right. In the case of @{term"gcd"} the
+rather than @{text"if"} on the right. In the case of @{term"gcd"} the
 following alternative definition suggests itself:
 *}
 
@@ -66,7 +67,7 @@
 @{prop"n ~= 0"}. Unfortunately, in general the case distinction
 may not be expressible by pattern matching.
 
-A very simple alternative is to replace @{name"if"} by @{name"case"}, which
+A very simple alternative is to replace @{text"if"} by @{text"case"}, which
 is also available for @{typ"bool"} but is not split automatically:
 *}