# HG changeset patch # User eberlm # Date 1491240736 -7200 # Node ID 65dd93a9f5b8ce0556be42b9368ae193f83f01fc # Parent b149abe619f76a7c1827ac0b981a7edf5eb8a481# Parent 6e47bcf7bec4b8217742ce0bee308379a364f1cf Merged diff -r b149abe619f7 -r 65dd93a9f5b8 src/Doc/Prog_Prove/Isar.thy --- a/src/Doc/Prog_Prove/Isar.thy Mon Apr 03 16:56:45 2017 +0200 +++ b/src/Doc/Prog_Prove/Isar.thy Mon Apr 03 19:32:16 2017 +0200 @@ -227,6 +227,10 @@ default. The patterns are phrased in terms of \isacom{show} but work for \isacom{have} and \isacom{lemma}, too. +\ifsem\else +\subsection{Logic} +\fi + We start with two forms of \concept{case analysis}: starting from a formula @{text P} we have the two cases @{text P} and @{prop"~P"}, and starting from a fact @{prop"P \ Q"} @@ -241,11 +245,11 @@ proof cases assume "P" text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "R" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show "R" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} next assume "\ P" text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "R" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show "R" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} qed(*<*)oops(*>*) text_raw {* } \end{minipage}\index{cases@@{text cases}} @@ -254,16 +258,16 @@ \isa{% *} (*<*)lemma "R" proof-(*>*) -have "P \ Q" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} +have "P \ Q" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} then show "R" proof assume "P" text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "R" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show "R" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} next assume "Q" text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "R" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show "R" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} qed(*<*)oops(*>*) text_raw {* } @@ -280,11 +284,11 @@ proof assume "P" text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "Q" (*<*)sorry(*>*) text_raw{*\ $\dots$\\*} + show "Q" (*<*)sorry(*>*) text_raw{*\ \isasymproof\\*} next assume "Q" text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "P" (*<*)sorry(*>*) text_raw{*\ $\dots$\\*} + show "P" (*<*)sorry(*>*) text_raw{*\ \isasymproof\\*} qed(*<*)qed(*>*) text_raw {* } \medskip @@ -300,7 +304,7 @@ proof assume "P" text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "False" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show "False" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} qed(*<*)oops(*>*) text_raw {* } @@ -314,7 +318,7 @@ proof (rule ccontr) assume "\P" text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "False" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show "False" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} qed(*<*)oops(*>*) text_raw {* } @@ -333,7 +337,7 @@ proof fix x text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "P(x)" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show "P(x)" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} qed(*<*)oops(*>*) text_raw {* } @@ -346,7 +350,7 @@ show "\x. P(x)" proof text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "P(witness)" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show "P(witness)" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} qed (*<*)oops(*>*) @@ -367,7 +371,7 @@ \end{isamarkuptext}% *} (*<*)lemma True proof- assume 1: "EX x. P x"(*>*) -have "\x. P(x)" (*<*)by(rule 1)(*>*)text_raw{*\ $\dots$\\*} +have "\x. P(x)" (*<*)by(rule 1)(*>*)text_raw{*\ \isasymproof\\*} then obtain x where p: "P(x)" by blast (*<*)oops(*>*) text{* @@ -401,9 +405,9 @@ (*<*)lemma "A = (B::'a set)" proof-(*>*) show "A = B" proof - show "A \ B" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show "A \ B" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} next - show "B \ A" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show "B \ A" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} qed(*<*)qed(*>*) text_raw {* } @@ -418,13 +422,92 @@ fix x assume "x \ A" text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "x \ B" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show "x \ B" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} qed(*<*)qed(*>*) text_raw {* } \end{minipage} \end{tabular} \begin{isamarkuptext}% + +\ifsem\else +\subsection{Chains of (In)Equations} + +In textbooks, chains of equations (and inequations) are often displayed like this: +\begin{quote} +\begin{tabular}{@ {}l@ {\qquad}l@ {}} +$t_1 = t_2$ & \isamath{\,\langle\mathit{justification}\rangle}\\ +$\phantom{t_1} = t_3$ & \isamath{\,\langle\mathit{justification}\rangle}\\ +\quad $\vdots$\\ +$\phantom{t_1} = t_n$ & \isamath{\,\langle\mathit{justification}\rangle} +\end{tabular} +\end{quote} +The Isar equivalent is this: + +\begin{samepage} +\begin{quote} +\isakeyword{have} \"t\<^sub>1 = t\<^sub>2"\ \isasymproof\\ +\isakeyword{also have} \"... = t\<^sub>3"\ \isasymproof\\ +\quad $\vdots$\\ +\isakeyword{also have} \"... = t\<^sub>n"\ \isasymproof \\ +\isakeyword{finally show} \"t\<^sub>1 = t\<^sub>n"\\ \texttt{.} +\end{quote} +\end{samepage} + +\noindent +The ``\...\'' and ``\.\'' deserve some explanation: +\begin{description} +\item[``\...\''] is literally three dots. It is the name of an unknown that Isar +automatically instantiates with the right-hand side of the previous equation. +In general, if \this\ is the theorem @{term "p t\<^sub>1 t\<^sub>2"} then ``\...\'' +stands for \t\<^sub>2\. +\item[``\.\''] (a single dot) is a proof method that solves a goal by one of the +assumptions. This works because the result of \isakeyword{finally} +is the theorem \mbox{\t\<^sub>1 = t\<^sub>n\}, +\isakeyword{show} \"t\<^sub>1 = t\<^sub>n"\ states the theorem explicitly, +and ``\.\'' proves the theorem with the result of \isakeyword{finally}. +\end{description} +The above proof template also works for arbitrary mixtures of \=\, \\\ and \<\, +for example: +\begin{quote} +\isakeyword{have} \"t\<^sub>1 < t\<^sub>2"\ \isasymproof\\ +\isakeyword{also have} \"... = t\<^sub>3"\ \isasymproof\\ +\quad $\vdots$\\ +\isakeyword{also have} \"... \ t\<^sub>n"\ \isasymproof \\ +\isakeyword{finally show} \"t\<^sub>1 < t\<^sub>n"\\ \texttt{.} +\end{quote} +The relation symbol in the \isakeyword{finally} step needs to be the most precise one +possible. In the example above, you must not write \t\<^sub>1 \ t\<^sub>n\ instead of \mbox{\t\<^sub>1 < t\<^sub>n\}. + +\begin{warn} +Isabelle only supports \=\, \\\ and \<\ but not \\\ and \>\ +in (in)equation chains (by default). +\end{warn} + +If you want to go beyond merely using the above proof patterns and want to +understand what \isakeyword{also} and \isakeyword{finally} mean, read on. +There is an Isar theorem variable called \calculation\, similar to \this\. +When the first \isakeyword{also} in a chain is encountered, Isabelle sets +\calculation := this\. In each subsequent \isakeyword{also} step, +Isabelle composes the theorems \calculation\ and \this\ (i.e.\ the two previous +(in)equalities) using some predefined set of rules including transitivity +of \=\, \\\ and \<\ but also mixed rules like @{prop"\ x \ y; y < z \ \ x < z"}. +The result of this composition is assigned to \calculation\. Consider +\begin{quote} +\isakeyword{have} \"t\<^sub>1 \ t\<^sub>2"\ \isasymproof\\ +\isakeyword{also} \isakeyword{have} \"... < t\<^sub>3"\ \isasymproof\\ +\isakeyword{also} \isakeyword{have} \"... = t\<^sub>4"\ \isasymproof\\ +\isakeyword{finally show} \"t\<^sub>1 < t\<^sub>4"\\ \texttt{.} +\end{quote} +After the first \isakeyword{also}, \calculation\ is \"t\<^sub>1 \ t\<^sub>2"\, +and after the second \isakeyword{also}, \calculation\ is \"t\<^sub>1 < t\<^sub>3"\. +The command \isakeyword{finally} is short for \isakeyword{also from} \calculation\. +Therefore the \isakeyword{also} hidden in \isakeyword{finally} sets \calculation\ +to \t\<^sub>1 < t\<^sub>4\ and the final ``\texttt{.}'' succeeds. + +For more information on this style of proof see \cite{BauerW-TPHOLs01}. +\fi + \section{Streamlining Proofs} \subsection{Pattern Matching and Quotations} @@ -445,11 +528,11 @@ proof assume "?L" text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "?R" (*<*)sorry(*>*) text_raw{*\ $\dots$\\*} + show "?R" (*<*)sorry(*>*) text_raw{*\ \isasymproof\\*} next assume "?R" text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show "?L" (*<*)sorry(*>*) text_raw{*\ $\dots$\\*} + show "?L" (*<*)sorry(*>*) text_raw{*\ \isasymproof\\*} qed(*<*)qed(*>*) text{* Instead of duplicating @{text"formula\<^sub>i"} in the text, we introduce @@ -463,7 +546,7 @@ lemma "formula" proof - text_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*} - show ?thesis (*<*)sorry(*>*) text_raw{*\ $\dots$\\*} + show ?thesis (*<*)sorry(*>*) text_raw{*\ \isasymproof\\*} qed text{* @@ -511,12 +594,12 @@ \isa{% *} (*<*)lemma "P" proof-(*>*) -have "P\<^sub>1" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} -moreover have "P\<^sub>2" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} +have "P\<^sub>1" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} +moreover have "P\<^sub>2" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} moreover text_raw{*\\$\vdots$\\\hspace{-1.4ex}*}(*<*)have "True" ..(*>*) -moreover have "P\<^sub>n" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} -ultimately have "P" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} +moreover have "P\<^sub>n" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} +ultimately have "P" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} (*<*)oops(*>*) text_raw {* } @@ -527,20 +610,20 @@ \isa{% *} (*<*)lemma "P" proof-(*>*) -have lab\<^sub>1: "P\<^sub>1" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} -have lab\<^sub>2: "P\<^sub>2" (*<*)sorry(*>*)text_raw{*\ $\dots$*} +have lab\<^sub>1: "P\<^sub>1" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} +have lab\<^sub>2: "P\<^sub>2" (*<*)sorry(*>*)text_raw{*\ \isasymproof*} text_raw{*\\$\vdots$\\\hspace{-1.4ex}*} -have lab\<^sub>n: "P\<^sub>n" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} +have lab\<^sub>n: "P\<^sub>n" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} from lab\<^sub>1 lab\<^sub>2 text_raw{*\ $\dots$\\*} -have "P" (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} +have "P" (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} (*<*)oops(*>*) text_raw {* } \end{minipage} \end{tabular} \begin{isamarkuptext}% -The \isacom{moreover} version is no shorter but expresses the structure more -clearly and avoids new names. +The \isacom{moreover} version is no shorter but expresses the structure +a bit more clearly and avoids new names. \subsection{Raw Proof Blocks} @@ -550,9 +633,9 @@ variables at the end. This is what a \concept{raw proof block} does: \begin{quote}\index{$IMP053@@{text"{ ... }"} (proof block)} @{text"{"} \isacom{fix} @{text"x\<^sub>1 \ x\<^sub>n"}\\ -\mbox{}\ \ \ \isacom{assume} @{text"A\<^sub>1 \ A\<^sub>m"}\\ -\mbox{}\ \ \ $\vdots$\\ -\mbox{}\ \ \ \isacom{have} @{text"B"}\\ +\mbox{}\ \ \isacom{assume} @{text"A\<^sub>1 \ A\<^sub>m"}\\ +\mbox{}\ \ $\vdots$\\ +\mbox{}\ \ \isacom{have} @{text"B"}\ \isasymproof\\ @{text"}"} \end{quote} proves @{text"\ A\<^sub>1; \ ; A\<^sub>m \ \ B"} @@ -734,11 +817,11 @@ proof (induction n) case 0 text_raw{*\\\mbox{}\ \ $\vdots$\\\mbox{}\hspace{-1ex}*} - show ?case (*<*)sorry(*>*) text_raw{*\ $\dots$\\*} + show ?case (*<*)sorry(*>*) text_raw{*\ \isasymproof\\*} next case (Suc n) text_raw{*\\\mbox{}\ \ $\vdots$\\\mbox{}\hspace{-1ex}*} - show ?case (*<*)sorry(*>*) text_raw{*\ $\dots$\\*} + show ?case (*<*)sorry(*>*) text_raw{*\ \isasymproof\\*} qed(*<*)qed(*>*) text_raw {* } @@ -905,7 +988,7 @@ proof(induction rule: I.induct) case rule\<^sub>1 text_raw{*\\[-.4ex]\mbox{}\ \ $\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}*} - show ?case (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show ?case (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} next text_raw{*\\[-.4ex]$\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}*} (*<*) @@ -915,7 +998,7 @@ next case rule\<^sub>n text_raw{*\\[-.4ex]\mbox{}\ \ $\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}*} - show ?case (*<*)sorry(*>*)text_raw{*\ $\dots$\\*} + show ?case (*<*)sorry(*>*)text_raw{*\ \isasymproof\\*} qed(*<*)qed(*>*) text{* diff -r b149abe619f7 -r 65dd93a9f5b8 src/Doc/Prog_Prove/document/root.bib --- a/src/Doc/Prog_Prove/document/root.bib Mon Apr 03 16:56:45 2017 +0200 +++ b/src/Doc/Prog_Prove/document/root.bib Mon Apr 03 19:32:16 2017 +0200 @@ -1,6 +1,12 @@ @string{CUP="Cambridge University Press"} @string{LNCS="Lect.\ Notes in Comp.\ Sci."} -@string{Springer="Springer-Verlag"} +@string{Springer="Springer"} + +@InProceedings{BauerW-TPHOLs01,author={Gertrud Bauer and Markus Wenzel}, +title={Calculational Reasoning Revisited --- An {Isabelle/Isar} Experience}, +booktitle={Theorem Proving in Higher Order Logics, TPHOLs 2001}, +editor={R. Boulton and P. Jackson}, +year=2001,publisher=Springer,series=LNCS,volume=2152,pages="75--90"} @book{HuthRyan,author="Michael Huth and Mark Ryan", title={Logic in Computer Science},publisher=CUP,year=2004}