diff -r 419116f1157a -r e23770bc97c8 doc-src/IsarRef/Thy/Proof.thy --- a/doc-src/IsarRef/Thy/Proof.thy Thu Feb 26 08:44:44 2009 -0800 +++ b/doc-src/IsarRef/Thy/Proof.thy Thu Feb 26 08:48:33 2009 -0800 @@ -1,17 +1,15 @@ -(* $Id$ *) - theory Proof imports Main begin -chapter {* Proofs *} +chapter {* Proofs \label{ch:proofs} *} text {* Proof commands perform transitions of Isar/VM machine configurations, which are block-structured, consisting of a stack of nodes with three main components: logical proof context, current - facts, and open goals. Isar/VM transitions are \emph{typed} - according to the following three different modes of operation: + facts, and open goals. Isar/VM transitions are typed according to + the following three different modes of operation: \begin{description} @@ -32,13 +30,17 @@ \end{description} - The proof mode indicator may be read as a verb telling the writer - what kind of operation may be performed next. The corresponding - typings of proof commands restricts the shape of well-formed proof - texts to particular command sequences. So dynamic arrangements of - commands eventually turn out as static texts of a certain structure. - \Appref{ap:refcard} gives a simplified grammar of the overall - (extensible) language emerging that way. + The proof mode indicator may be understood as an instruction to the + writer, telling what kind of operation may be performed next. The + corresponding typings of proof commands restricts the shape of + well-formed proof texts to particular command sequences. So dynamic + arrangements of commands eventually turn out as static texts of a + certain structure. + + \Appref{ap:refcard} gives a simplified grammar of the (extensible) + language emerging that way from the different types of proof + commands. The main ideas of the overall Isar framework are + explained in \chref{ch:isar-framework}. *} @@ -963,7 +965,7 @@ \begin{matharray}{l} @{text "\using b\<^sub>1 \ b\<^sub>k\"}~~@{command "obtain"}~@{text "x\<^sub>1 \ x\<^sub>m \ a: \\<^sub>1 \ \\<^sub>n \proof\ \"} \\[1ex] \quad @{command "have"}~@{text "\thesis. (\x\<^sub>1 \ x\<^sub>m. \\<^sub>1 \ \ \\<^sub>n \ thesis) \ thesis"} \\ - \quad @{command "proof"}~@{text succeed} \\ + \quad @{command "proof"}~@{method succeed} \\ \qquad @{command "fix"}~@{text thesis} \\ \qquad @{command "assume"}~@{text "that [Pure.intro?]: \x\<^sub>1 \ x\<^sub>m. \\<^sub>1 \ \ \\<^sub>n \ thesis"} \\ \qquad @{command "then"}~@{command "show"}~@{text thesis} \\