diff -r a2bff98d6e5d -r 9710486b886b doc-src/TutorialI/Misc/AdvancedInd.thy --- a/doc-src/TutorialI/Misc/AdvancedInd.thy Tue May 01 22:26:55 2001 +0200 +++ b/doc-src/TutorialI/Misc/AdvancedInd.thy Wed May 02 11:54:18 2001 +0200 @@ -111,11 +111,14 @@ intuitive. For example, applying wellfounded induction on $x$ (w.r.t.\ $\prec$) to (\ref{eqn:ind-over-term}) and transforming the result a little leads to the goal -\[ \bigwedge\vec{y}.~\forall \vec{z}.~ t(\vec{z}) \prec t(\vec{y}) \longrightarrow C(\vec{z}) - \Longrightarrow C(\vec{y}) \] -where the dependence of $t$ and $C$ on their free variables has been made explicit. -Unfortunately, this induction schema cannot be expressed as a single theorem because it depends -on the number of free variables in $t$ --- the notation $\vec{y}$ is merely an informal device. +\[ \bigwedge\overline{y}.\ + \forall \overline{z}.\ t\,\overline{z} \prec t\,\overline{y}\ \longrightarrow\ C\,\overline{z} + \ \Longrightarrow\ C\,\overline{y} \] +where $\overline{y}$ stands for $y@1 \dots y@n$ and the dependence of $t$ and +$C$ on the free variables of $t$ has been made explicit. +Unfortunately, this induction schema cannot be expressed as a +single theorem because it depends on the number of free variables in $t$ --- +the notation $\overline{y}$ is merely an informal device. *} subsection{*Beyond Structural and Recursion Induction*}; @@ -133,7 +136,8 @@ induction, where you must prove $P(n)$ under the assumption that $P(m)$ holds for all $m nat";