diff -r bda7527ccf05 -r cc2d676d5395 src/Doc/Tutorial/Misc/natsum.thy --- a/src/Doc/Tutorial/Misc/natsum.thy Wed Dec 26 16:07:28 2018 +0100 +++ b/src/Doc/Tutorial/Misc/natsum.thy Wed Dec 26 16:25:20 2018 +0100 @@ -2,7 +2,7 @@ theory natsum imports Main begin (*>*) text\\noindent -In particular, there are @{text"case"}-expressions, for example +In particular, there are \case\-expressions, for example @{term[display]"case n of 0 => 0 | Suc m => m"} primitive recursion, for example \ @@ -38,17 +38,17 @@ \cdx{max}, \isadxboldpos{\isasymle}{$HOL2arithrel} and \isadxboldpos{<}{$HOL2arithrel} are overloaded: they are available not just for natural numbers but for other types as well. - For example, given the goal @{text"x + 0 = x"}, there is nothing to indicate + For example, given the goal \x + 0 = x\, there is nothing to indicate that you are talking about natural numbers. Hence Isabelle can only infer - that @{term x} is of some arbitrary type where @{text 0} and @{text"+"} are + that @{term x} is of some arbitrary type where \0\ and \+\ are declared. As a consequence, you will be unable to prove the goal. To alert you to such pitfalls, Isabelle flags numerals without a fixed type in its output: @{prop"x+0 = x"}. (In the absence of a numeral, it may take you some time to realize what has happened if \pgmenu{Show Types} is not set). In this particular example, you need to include - an explicit type constraint, for example @{text"x+0 = (x::nat)"}. If there + an explicit type constraint, for example \x+0 = (x::nat)\. If there is enough contextual information this may not be necessary: @{prop"Suc x = - x"} automatically implies @{text"x::nat"} because @{term Suc} is not + x"} automatically implies \x::nat\ because @{term Suc} is not overloaded. For details on overloading see \S\ref{sec:overloading}. @@ -57,20 +57,20 @@ \end{warn} \begin{warn} The symbols \isadxboldpos{>}{$HOL2arithrel} and - \isadxboldpos{\isasymge}{$HOL2arithrel} are merely syntax: @{text"x > y"} - stands for @{prop"y < x"} and similary for @{text"\"} and - @{text"\"}. + \isadxboldpos{\isasymge}{$HOL2arithrel} are merely syntax: \x > y\ + stands for @{prop"y < x"} and similary for \\\ and + \\\. \end{warn} \begin{warn} - Constant @{text"1::nat"} is defined to equal @{term"Suc 0"}. This definition + Constant \1::nat\ is defined to equal @{term"Suc 0"}. This definition (see \S\ref{sec:ConstDefinitions}) is unfolded automatically by some - tactics (like @{text auto}, @{text simp} and @{text arith}) but not by + tactics (like \auto\, \simp\ and \arith\) but not by others (especially the single step tactics in Chapter~\ref{chap:rules}). If you need the full set of numerals, see~\S\ref{sec:numerals}. \emph{Novices are advised to stick to @{term"0::nat"} and @{term Suc}.} \end{warn} -Both @{text auto} and @{text simp} +Both \auto\ and \simp\ (a method introduced below, \S\ref{sec:Simplification}) prove simple arithmetic goals automatically: \ @@ -81,7 +81,7 @@ text\\noindent For efficiency's sake, this built-in prover ignores quantified formulae, many logical connectives, and all arithmetic operations apart from addition. -In consequence, @{text auto} and @{text simp} cannot prove this slightly more complex goal: +In consequence, \auto\ and \simp\ cannot prove this slightly more complex goal: \ lemma "m \ (n::nat) \ m < n \ n < m" @@ -89,10 +89,10 @@ text\\noindent The method \methdx{arith} is more general. It attempts to prove the first subgoal provided it is a \textbf{linear arithmetic} formula. -Such formulas may involve the usual logical connectives (@{text"\"}, -@{text"\"}, @{text"\"}, @{text"\"}, @{text"="}, -@{text"\"}, @{text"\"}), the relations @{text"="}, -@{text"\"} and @{text"<"}, and the operations @{text"+"}, @{text"-"}, +Such formulas may involve the usual logical connectives (\\\, +\\\, \\\, \\\, \=\, +\\\, \\\), the relations \=\, +\\\ and \<\, and the operations \+\, \-\, @{term min} and @{term max}. For example,\ lemma "min i (max j (k*k)) = max (min (k*k) i) (min i (j::nat))" @@ -107,19 +107,19 @@ (*<*)oops(*>*) text\\noindent -is not proved by @{text arith} because the proof relies +is not proved by \arith\ because the proof relies on properties of multiplication. Only multiplication by numerals (which is the same as iterated addition) is taken into account. -\begin{warn} The running time of @{text arith} is exponential in the number +\begin{warn} The running time of \arith\ is exponential in the number of occurrences of \ttindexboldpos{-}{$HOL2arithfun}, \cdx{min} and \cdx{max} because they are first eliminated by case distinctions. -If @{text k} is a numeral, \sdx{div}~@{text k}, \sdx{mod}~@{text k} and -@{text k}~\sdx{dvd} are also supported, where the former two are eliminated +If \k\ is a numeral, \sdx{div}~\k\, \sdx{mod}~\k\ and +\k\~\sdx{dvd} are also supported, where the former two are eliminated by case distinctions, again blowing up the running time. -If the formula involves quantifiers, @{text arith} may take +If the formula involves quantifiers, \arith\ may take super-exponential time and space. \end{warn} \