diff -r bda7527ccf05 -r cc2d676d5395 src/Doc/Tutorial/Misc/AdvancedInd.thy --- a/src/Doc/Tutorial/Misc/AdvancedInd.thy Wed Dec 26 16:07:28 2018 +0100 +++ b/src/Doc/Tutorial/Misc/AdvancedInd.thy Wed Dec 26 16:25:20 2018 +0100 @@ -45,12 +45,12 @@ heuristic applies to rule inductions; see \S\ref{sec:rtc}.} \begin{quote} \emph{Pull all occurrences of the induction variable into the conclusion -using @{text"\"}.} +using \\\.} \end{quote} Thus we should state the lemma as an ordinary -implication~(@{text"\"}), letting +implication~(\\\), letting \attrdx{rule_format} (\S\ref{sec:forward}) convert the -result to the usual @{text"\"} form: +result to the usual \\\ form: \ (*<*)oops(*>*) lemma hd_rev [rule_format]: "xs \ [] \ hd(rev xs) = last xs" @@ -61,15 +61,15 @@ txt\\noindent This time, induction leaves us with a trivial base case: @{subgoals[display,indent=0,goals_limit=1]} -And @{text"auto"} completes the proof. +And \auto\ completes the proof. If there are multiple premises $A@1$, \dots, $A@n$ containing the induction variable, you should turn the conclusion $C$ into \[ A@1 \longrightarrow \cdots A@n \longrightarrow C. \] Additionally, you may also have to universally quantify some other variables, -which can yield a fairly complex conclusion. However, @{text rule_format} -can remove any number of occurrences of @{text"\"} and -@{text"\"}. +which can yield a fairly complex conclusion. However, \rule_format\ +can remove any number of occurrences of \\\ and +\\\. \index{induction!on a term}% A second reason why your proposition may not be amenable to induction is that @@ -92,7 +92,7 @@ For an example see \S\ref{sec:CTL-revisited} below. Of course, all premises that share free variables with $t$ need to be pulled into -the conclusion as well, under the @{text"\"}, again using @{text"\"} as shown above. +the conclusion as well, under the \\\, again using \\\ as shown above. Readers who are puzzled by the form of statement (\ref{eqn:ind-over-term}) above should remember that the @@ -138,7 +138,7 @@ text\ \begin{warn} We discourage the use of axioms because of the danger of -inconsistencies. Axiom @{text f_ax} does +inconsistencies. Axiom \f_ax\ does not introduce an inconsistency because, for example, the identity function satisfies it. Axioms can be useful in exploratory developments, say when you assume some well-known theorems so that you can quickly demonstrate some @@ -163,7 +163,7 @@ txt\\noindent We get the following proof state: @{subgoals[display,indent=0,margin=65]} -After stripping the @{text"\i"}, the proof continues with a case +After stripping the \\i\, the proof continues with a case distinction on @{term"i"}. The case @{prop"i = (0::nat)"} is trivial and we focus on the other case: \ @@ -208,7 +208,7 @@ lemmas f_incr = f_incr_lem[rule_format, OF refl] text\\noindent -The final @{thm[source]refl} gets rid of the premise @{text"?k = f ?i"}. +The final @{thm[source]refl} gets rid of the premise \?k = f ?i\. We could have included this derivation in the original statement of the lemma: \ @@ -225,7 +225,7 @@ whose conclusion is of the form ${?}P~?x@1 \dots ?x@n$, in which case the format is \begin{quote} -\isacommand{apply}@{text"(induct_tac"} $y@1 \dots y@n$ @{text"rule:"} $r$@{text")"} +\isacommand{apply}\(induct_tac\ $y@1 \dots y@n$ \rule:\ $r$\)\ \end{quote} where $y@1, \dots, y@n$ are variables in the conclusion of the first subgoal. @@ -279,7 +279,7 @@ HOL already provides the mother of all inductions, well-founded induction (see \S\ref{sec:Well-founded}). For example theorem @{thm[source]nat_less_induct} is -a special case of @{thm[source]wf_induct} where @{term r} is @{text"<"} on +a special case of @{thm[source]wf_induct} where @{term r} is \<\ on @{typ nat}. The details can be found in theory \isa{Wellfounded_Recursion}. \ (*<*)end(*>*)