# HG changeset patch # User wenzelm # Date 1307521447 -7200 # Node ID bc72c1ccc89e402d4dcf2a2f20085c1f3236f74d # Parent 3535f16d97145c40bc37f4891fc68cc85bead9c8# Parent 48a0a9db3453edc587fd30f7509d34424f2cd1ff merged diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/HOL/HOL.tex --- a/doc-src/HOL/HOL.tex Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/HOL/HOL.tex Wed Jun 08 10:24:07 2011 +0200 @@ -2,38 +2,6 @@ \index{higher-order logic|(} \index{HOL system@{\sc hol} system} -The theory~\thydx{HOL} implements higher-order logic. It is based on -Gordon's~{\sc hol} system~\cite{mgordon-hol}, which itself is based on -Church's original paper~\cite{church40}. Andrews's book~\cite{andrews86} is a -full description of the original Church-style higher-order logic. Experience -with the {\sc hol} system has demonstrated that higher-order logic is widely -applicable in many areas of mathematics and computer science, not just -hardware verification, {\sc hol}'s original \textit{raison d'{\^e}tre\/}. It -is weaker than ZF set theory but for most applications this does not matter. -If you prefer {\ML} to Lisp, you will probably prefer HOL to~ZF. - -The syntax of HOL\footnote{Earlier versions of Isabelle's HOL used a different - syntax. Ancient releases of Isabelle included still another version of~HOL, - with explicit type inference rules~\cite{paulson-COLOG}. This version no - longer exists, but \thydx{ZF} supports a similar style of reasoning.} -follows $\lambda$-calculus and functional programming. Function application -is curried. To apply the function~$f$ of type $\tau@1\To\tau@2\To\tau@3$ to -the arguments~$a$ and~$b$ in HOL, you simply write $f\,a\,b$. There is no -`apply' operator as in \thydx{ZF}. Note that $f(a,b)$ means ``$f$ applied to -the pair $(a,b)$'' in HOL. We write ordered pairs as $(a,b)$, not $\langle -a,b\rangle$ as in ZF. - -HOL has a distinct feel, compared with ZF and CTT. It identifies object-level -types with meta-level types, taking advantage of Isabelle's built-in -type-checker. It identifies object-level functions with meta-level functions, -so it uses Isabelle's operations for abstraction and application. - -These identifications allow Isabelle to support HOL particularly nicely, but -they also mean that HOL requires more sophistication from the user --- in -particular, an understanding of Isabelle's type system. Beginners should work -with \texttt{show_types} (or even \texttt{show_sorts}) set to \texttt{true}. - - \begin{figure} \begin{constants} \it name &\it meta-type & \it description \\ @@ -147,9 +115,9 @@ term} if $\sigma$ and~$\tau$ do, allowing quantification over functions. -HOL allows new types to be declared as subsets of existing types; -see~{\S}\ref{sec:HOL:Types}. ML-like datatypes can also be declared; see -~{\S}\ref{sec:HOL:datatype}. +HOL allows new types to be declared as subsets of existing types, +either using the primitive \texttt{typedef} or the more convenient +\texttt{datatype} (see~{\S}\ref{sec:HOL:datatype}). Several syntactic type classes --- \cldx{plus}, \cldx{minus}, \cldx{times} and @@ -342,8 +310,7 @@ \begin{warn} The definitions above should never be expanded and are shown for completeness only. Instead users should reason in terms of the derived rules shown below -or, better still, using high-level tactics -(see~{\S}\ref{sec:HOL:generic-packages}). +or, better still, using high-level tactics. \end{warn} Some of the rules mention type variables; for example, \texttt{refl} @@ -912,13 +879,7 @@ on sets in the file \texttt{HOL/mono.ML}. -\section{Generic packages} -\label{sec:HOL:generic-packages} - -HOL instantiates most of Isabelle's generic packages, making available the -simplifier and the classical reasoner. - -\subsection{Simplification and substitution} +\section{Simplification and substitution} Simplification tactics tactics such as \texttt{Asm_simp_tac} and \texttt{Full_simp_tac} use the default simpset (\texttt{simpset()}), which works for most purposes. A quite minimal @@ -964,7 +925,7 @@ equality throughout a subgoal and its hypotheses. This tactic uses HOL's general substitution rule. -\subsubsection{Case splitting} +\subsection{Case splitting} \label{subsec:HOL:case:splitting} HOL also provides convenient means for case splitting during rewriting. Goals @@ -1021,115 +982,6 @@ \end{ttbox} for adding splitting rules to, and deleting them from the current simpset. -\subsection{Classical reasoning} - -HOL derives classical introduction rules for $\disj$ and~$\exists$, as well as -classical elimination rules for~$\imp$ and~$\bimp$, and the swap rule; recall -Fig.\ts\ref{hol-lemmas2} above. - -The classical reasoner is installed. Tactics such as \texttt{Blast_tac} and -{\tt Best_tac} refer to the default claset (\texttt{claset()}), which works -for most purposes. Named clasets include \ttindexbold{prop_cs}, which -includes the propositional rules, and \ttindexbold{HOL_cs}, which also -includes quantifier rules. See the file \texttt{HOL/cladata.ML} for lists of -the classical rules, -and \iflabelundefined{chap:classical}{the {\em Reference Manual\/}}% -{Chap.\ts\ref{chap:classical}} for more discussion of classical proof methods. - - -%FIXME outdated, both from the Isabelle and SVC perspective -% \section{Calling the decision procedure SVC}\label{sec:HOL:SVC} - -% \index{SVC decision procedure|(} - -% The Stanford Validity Checker (SVC) is a tool that can check the validity of -% certain types of formulae. If it is installed on your machine, then -% Isabelle/HOL can be configured to call it through the tactic -% \ttindex{svc_tac}. It is ideal for large tautologies and complex problems in -% linear arithmetic. Subexpressions that SVC cannot handle are automatically -% replaced by variables, so you can call the tactic on any subgoal. See the -% file \texttt{HOL/ex/svc_test.ML} for examples. -% \begin{ttbox} -% svc_tac : int -> tactic -% Svc.trace : bool ref \hfill{\bf initially false} -% \end{ttbox} - -% \begin{ttdescription} -% \item[\ttindexbold{svc_tac} $i$] attempts to prove subgoal~$i$ by translating -% it into a formula recognized by~SVC\@. If it succeeds then the subgoal is -% removed. It fails if SVC is unable to prove the subgoal. It crashes with -% an error message if SVC appears not to be installed. Numeric variables may -% have types \texttt{nat}, \texttt{int} or \texttt{real}. - -% \item[\ttindexbold{Svc.trace}] is a flag that, if set, causes \texttt{svc_tac} -% to trace its operations: abstraction of the subgoal, translation to SVC -% syntax, SVC's response. -% \end{ttdescription} - -% Here is an example, with tracing turned on: -% \begin{ttbox} -% set Svc.trace; -% {\out val it : bool = true} -% Goal "(#3::nat)*a <= #2 + #4*b + #6*c & #11 <= #2*a + b + #2*c & \ttback -% \ttback a + #3*b <= #5 + #2*c --> #2 + #3*b <= #2*a + #6*c"; - -% by (svc_tac 1); -% {\out Subgoal abstracted to} -% {\out #3 * a <= #2 + #4 * b + #6 * c &} -% {\out #11 <= #2 * a + b + #2 * c & a + #3 * b <= #5 + #2 * c -->} -% {\out #2 + #3 * b <= #2 * a + #6 * c} -% {\out Calling SVC:} -% {\out (=> (<= 0 (F_c) ) (=> (<= 0 (F_b) ) (=> (<= 0 (F_a) )} -% {\out (=> (AND (<= {* 3 (F_a) } {+ {+ 2 {* 4 (F_b) } } } -% {\out {* 6 (F_c) } } ) (AND (<= 11 {+ {+ {* 2 (F_a) } (F_b) }} -% {\out {* 2 (F_c) } } ) (<= {+ (F_a) {* 3 (F_b) } } {+ 5 } -% {\out {* 2 (F_c) } } ) ) ) (< {+ 2 {* 3 (F_b) } } {+ 1 {+ } -% {\out {* 2 (F_a) } {* 6 (F_c) } } } ) ) ) ) ) } -% {\out SVC Returns:} -% {\out VALID} -% {\out Level 1} -% {\out #3 * a <= #2 + #4 * b + #6 * c &} -% {\out #11 <= #2 * a + b + #2 * c & a + #3 * b <= #5 + #2 * c -->} -% {\out #2 + #3 * b <= #2 * a + #6 * c} -% {\out No subgoals!} -% \end{ttbox} - - -% \begin{warn} -% Calling \ttindex{svc_tac} entails an above-average risk of -% unsoundness. Isabelle does not check SVC's result independently. Moreover, -% the tactic translates the submitted formula using code that lies outside -% Isabelle's inference core. Theorems that depend upon results proved using SVC -% (and other oracles) are displayed with the annotation \texttt{[!]} attached. -% You can also use \texttt{\#der (rep_thm $th$)} to examine the proof object of -% theorem~$th$, as described in the \emph{Reference Manual}. -% \end{warn} - -% To start, first download SVC from the Internet at URL -% \begin{ttbox} -% http://agamemnon.stanford.edu/~levitt/vc/index.html -% \end{ttbox} -% and install it using the instructions supplied. SVC requires two environment -% variables: -% \begin{ttdescription} -% \item[\ttindexbold{SVC_HOME}] is an absolute pathname to the SVC -% distribution directory. - -% \item[\ttindexbold{SVC_MACHINE}] identifies the type of computer and -% operating system. -% \end{ttdescription} -% You can set these environment variables either using the Unix shell or through -% an Isabelle \texttt{settings} file. Isabelle assumes SVC to be installed if -% \texttt{SVC_HOME} is defined. - -% \paragraph*{Acknowledgement.} This interface uses code supplied by S{\o}ren -% Heilmann. - - -% \index{SVC decision procedure|)} - - - \section{Types}\label{sec:HOL:Types} This section describes HOL's basic predefined types ($\alpha \times \beta$, @@ -1249,9 +1101,7 @@ shown. The constructions are fairly standard and can be found in the respective theory files. Although the sum and product types are constructed manually for foundational reasons, they are represented as -actual datatypes later (see {\S}\ref{subsec:datatype:rep_datatype}). -Therefore, the theory \texttt{Datatype} should be used instead of -\texttt{Sum} or \texttt{Prod}. +actual datatypes later. \begin{figure} \begin{constants} @@ -1375,7 +1225,7 @@ the order of the two cases. Both \texttt{primrec} and \texttt{case} are realized by a recursion operator \cdx{nat_rec}, which is available because \textit{nat} is represented as -a datatype (see {\S}\ref{subsec:datatype:rep_datatype}). +a datatype. %The predecessor relation, \cdx{pred_nat}, is shown to be well-founded. %Recursion along this relation resembles primitive recursion, but is @@ -1591,101 +1441,19 @@ \texttt{split_if} (see~{\S}\ref{subsec:HOL:case:splitting}). \texttt{List} provides a basic library of list processing functions defined by -primitive recursion (see~{\S}\ref{sec:HOL:primrec}). The recursion equations +primitive recursion. The recursion equations are shown in Figs.\ts\ref{fig:HOL:list-simps} and~\ref{fig:HOL:list-simps2}. \index{list@{\textit{list}} type|)} -\subsection{Introducing new types} \label{sec:typedef} - -The HOL-methodology dictates that all extensions to a theory should be -\textbf{definitional}. The type definition mechanism that meets this -criterion is \ttindex{typedef}. Note that \emph{type synonyms}, which are -inherited from Pure and described elsewhere, are just syntactic abbreviations -that have no logical meaning. - -\begin{warn} - Types in HOL must be non-empty; otherwise the quantifier rules would be - unsound, because $\exists x. x=x$ is a theorem \cite[{\S}7]{paulson-COLOG}. -\end{warn} -A \bfindex{type definition} identifies the new type with a subset of -an existing type. More precisely, the new type is defined by -exhibiting an existing type~$\tau$, a set~$A::\tau\,set$, and a -theorem of the form $x:A$. Thus~$A$ is a non-empty subset of~$\tau$, -and the new type denotes this subset. New functions are defined that -establish an isomorphism between the new type and the subset. If -type~$\tau$ involves type variables $\alpha@1$, \ldots, $\alpha@n$, -then the type definition creates a type constructor -$(\alpha@1,\ldots,\alpha@n)ty$ rather than a particular type. - -The syntax for type definitions is given in the Isabelle/Isar -reference manual. - -If all context conditions are met (no duplicate type variables in -`typevarlist', no extra type variables in `set', and no free term variables -in `set'), the following components are added to the theory: -\begin{itemize} -\item a type $ty :: (term,\dots,term)term$ -\item constants -\begin{eqnarray*} -T &::& \tau\;set \\ -Rep_T &::& (\alpha@1,\dots,\alpha@n)ty \To \tau \\ -Abs_T &::& \tau \To (\alpha@1,\dots,\alpha@n)ty -\end{eqnarray*} -\item a definition and three axioms -\[ -\begin{array}{ll} -T{\tt_def} & T \equiv A \\ -{\tt Rep_}T & Rep_T\,x \in T \\ -{\tt Rep_}T{\tt_inverse} & Abs_T\,(Rep_T\,x) = x \\ -{\tt Abs_}T{\tt_inverse} & y \in T \Imp Rep_T\,(Abs_T\,y) = y -\end{array} -\] -stating that $(\alpha@1,\dots,\alpha@n)ty$ is isomorphic to $A$ by $Rep_T$ -and its inverse $Abs_T$. -\end{itemize} -Below are two simple examples of HOL type definitions. Non-emptiness is -proved automatically here. -\begin{ttbox} -typedef unit = "{\ttlbrace}True{\ttrbrace}" - -typedef (prod) - ('a, 'b) "*" (infixr 20) - = "{\ttlbrace}f . EX (a::'a) (b::'b). f = (\%x y. x = a & y = b){\ttrbrace}" -\end{ttbox} - -Type definitions permit the introduction of abstract data types in a safe -way, namely by providing models based on already existing types. Given some -abstract axiomatic description $P$ of a type, this involves two steps: -\begin{enumerate} -\item Find an appropriate type $\tau$ and subset $A$ which has the desired - properties $P$, and make a type definition based on this representation. -\item Prove that $P$ holds for $ty$ by lifting $P$ from the representation. -\end{enumerate} -You can now forget about the representation and work solely in terms of the -abstract properties $P$. - -\begin{warn} -If you introduce a new type (constructor) $ty$ axiomatically, i.e.\ by -declaring the type and its operations and by stating the desired axioms, you -should make sure the type has a non-empty model. You must also have a clause -\par -\begin{ttbox} -arities \(ty\) :: (term,\thinspace\(\dots\),{\thinspace}term){\thinspace}term -\end{ttbox} -in your theory file to tell Isabelle that $ty$ is in class \texttt{term}, the -class of all HOL types. -\end{warn} - - \section{Datatype definitions} \label{sec:HOL:datatype} \index{*datatype|(} Inductive datatypes, similar to those of \ML, frequently appear in applications of Isabelle/HOL. In principle, such types could be defined by -hand via \texttt{typedef} (see {\S}\ref{sec:typedef}), but this would be far too +hand via \texttt{typedef}, but this would be far too tedious. The \ttindex{datatype} definition package of Isabelle/HOL (cf.\ \cite{Berghofer-Wenzel:1999:TPHOL}) automates such chores. It generates an appropriate \texttt{typedef} based on a least fixed-point construction, and @@ -2022,352 +1790,21 @@ and \texttt{simps} are contained in a separate structure named $t@1_\ldots_t@n$. -\subsection{Representing existing types as datatypes}\label{subsec:datatype:rep_datatype} - -For foundational reasons, some basic types such as \texttt{nat}, \texttt{*}, {\tt - +}, \texttt{bool} and \texttt{unit} are not defined in a \texttt{datatype} section, -but by more primitive means using \texttt{typedef}. To be able to use the -tactics \texttt{induct_tac} and \texttt{case_tac} and to define functions by -primitive recursion on these types, such types may be represented as actual -datatypes. This is done by specifying the constructors of the desired type, -plus a proof of the induction rule, as well as theorems -stating the distinctness and injectivity of constructors in a {\tt -rep_datatype} section. For the sum type this works as follows: -\begin{ttbox} -rep_datatype (sum) Inl Inr -proof - - fix P - fix s :: "'a + 'b" - assume x: "!!x::'a. P (Inl x)" and y: "!!y::'b. P (Inr y)" - then show "P s" by (auto intro: sumE [of s]) -qed simp_all -\end{ttbox} -The datatype package automatically derives additional theorems for recursion -and case combinators from these rules. Any of the basic HOL types mentioned -above are represented as datatypes. Try an induction on \texttt{bool} -today. - - -\subsection{Examples} - -\subsubsection{The datatype $\alpha~mylist$} - -We want to define a type $\alpha~mylist$. To do this we have to build a new -theory that contains the type definition. We start from the theory -\texttt{Datatype} instead of \texttt{Main} in order to avoid clashes with the -\texttt{List} theory of Isabelle/HOL. -\begin{ttbox} -MyList = Datatype + - datatype 'a mylist = Nil | Cons 'a ('a mylist) -end -\end{ttbox} -After loading the theory, we can prove $Cons~x~xs\neq xs$, for example. To -ease the induction applied below, we state the goal with $x$ quantified at the -object-level. This will be stripped later using \ttindex{qed_spec_mp}. -\begin{ttbox} -Goal "!x. Cons x xs ~= xs"; -{\out Level 0} -{\out ! x. Cons x xs ~= xs} -{\out 1. ! x. Cons x xs ~= xs} -\end{ttbox} -This can be proved by the structural induction tactic: -\begin{ttbox} -by (induct_tac "xs" 1); -{\out Level 1} -{\out ! x. Cons x xs ~= xs} -{\out 1. ! x. Cons x Nil ~= Nil} -{\out 2. !!a mylist.} -{\out ! x. Cons x mylist ~= mylist ==>} -{\out ! x. Cons x (Cons a mylist) ~= Cons a mylist} -\end{ttbox} -The first subgoal can be proved using the simplifier. Isabelle/HOL has -already added the freeness properties of lists to the default simplification -set. -\begin{ttbox} -by (Simp_tac 1); -{\out Level 2} -{\out ! x. Cons x xs ~= xs} -{\out 1. !!a mylist.} -{\out ! x. Cons x mylist ~= mylist ==>} -{\out ! x. Cons x (Cons a mylist) ~= Cons a mylist} -\end{ttbox} -Similarly, we prove the remaining goal. -\begin{ttbox} -by (Asm_simp_tac 1); -{\out Level 3} -{\out ! x. Cons x xs ~= xs} -{\out No subgoals!} -\ttbreak -qed_spec_mp "not_Cons_self"; -{\out val not_Cons_self = "Cons x xs ~= xs" : thm} -\end{ttbox} -Because both subgoals could have been proved by \texttt{Asm_simp_tac} -we could have done that in one step: -\begin{ttbox} -by (ALLGOALS Asm_simp_tac); -\end{ttbox} - - -\subsubsection{The datatype $\alpha~mylist$ with mixfix syntax} - -In this example we define the type $\alpha~mylist$ again but this time -we want to write \texttt{[]} for \texttt{Nil} and we want to use infix -notation \verb|#| for \texttt{Cons}. To do this we simply add mixfix -annotations after the constructor declarations as follows: -\begin{ttbox} -MyList = Datatype + - datatype 'a mylist = - Nil ("[]") | - Cons 'a ('a mylist) (infixr "#" 70) -end -\end{ttbox} -Now the theorem in the previous example can be written \verb|x#xs ~= xs|. - - -\subsubsection{A datatype for weekdays} - -This example shows a datatype that consists of 7 constructors: -\begin{ttbox} -Days = Main + - datatype days = Mon | Tue | Wed | Thu | Fri | Sat | Sun -end -\end{ttbox} -Because there are more than 6 constructors, inequality is expressed via a function -\verb|days_ord|. The theorem \verb|Mon ~= Tue| is not directly -contained among the distinctness theorems, but the simplifier can -prove it thanks to rewrite rules inherited from theory \texttt{NatArith}: -\begin{ttbox} -Goal "Mon ~= Tue"; -by (Simp_tac 1); -\end{ttbox} -You need not derive such inequalities explicitly: the simplifier will dispose -of them automatically. -\index{*datatype|)} - - -\section{Recursive function definitions}\label{sec:HOL:recursive} -\index{recursive functions|see{recursion}} - -Isabelle/HOL provides two main mechanisms of defining recursive functions. -\begin{enumerate} -\item \textbf{Primitive recursion} is available only for datatypes, and it is - somewhat restrictive. Recursive calls are only allowed on the argument's - immediate constituents. On the other hand, it is the form of recursion most - often wanted, and it is easy to use. - -\item \textbf{Well-founded recursion} requires that you supply a well-founded - relation that governs the recursion. Recursive calls are only allowed if - they make the argument decrease under the relation. Complicated recursion - forms, such as nested recursion, can be dealt with. Termination can even be - proved at a later time, though having unsolved termination conditions around - can make work difficult.% - \footnote{This facility is based on Konrad Slind's TFL - package~\cite{slind-tfl}. Thanks are due to Konrad for implementing TFL - and assisting with its installation.} -\end{enumerate} - -Following good HOL tradition, these declarations do not assert arbitrary -axioms. Instead, they define the function using a recursion operator. Both -HOL and ZF derive the theory of well-founded recursion from first -principles~\cite{paulson-set-II}. Primitive recursion over some datatype -relies on the recursion operator provided by the datatype package. With -either form of function definition, Isabelle proves the desired recursion -equations as theorems. - - -\subsection{Primitive recursive functions} -\label{sec:HOL:primrec} -\index{recursion!primitive|(} -\index{*primrec|(} - -Datatypes come with a uniform way of defining functions, {\bf primitive - recursion}. In principle, one could introduce primitive recursive functions -by asserting their reduction rules as new axioms, but this is not recommended: -\begin{ttbox}\slshape -Append = Main + -consts app :: ['a list, 'a list] => 'a list -rules - app_Nil "app [] ys = ys" - app_Cons "app (x#xs) ys = x#app xs ys" -end -\end{ttbox} -Asserting axioms brings the danger of accidentally asserting nonsense, as -in \verb$app [] ys = us$. - -The \ttindex{primrec} declaration is a safe means of defining primitive -recursive functions on datatypes: -\begin{ttbox} -Append = Main + -consts app :: ['a list, 'a list] => 'a list -primrec - "app [] ys = ys" - "app (x#xs) ys = x#app xs ys" -end -\end{ttbox} -Isabelle will now check that the two rules do indeed form a primitive -recursive definition. For example -\begin{ttbox} -primrec - "app [] ys = us" -\end{ttbox} -is rejected with an error message ``\texttt{Extra variables on rhs}''. - -\bigskip - -The general form of a primitive recursive definition is -\begin{ttbox} -primrec - {\it reduction rules} -\end{ttbox} -where \textit{reduction rules} specify one or more equations of the form -\[ f \, x@1 \, \dots \, x@m \, (C \, y@1 \, \dots \, y@k) \, z@1 \, -\dots \, z@n = r \] such that $C$ is a constructor of the datatype, $r$ -contains only the free variables on the left-hand side, and all recursive -calls in $r$ are of the form $f \, \dots \, y@i \, \dots$ for some $i$. There -must be at most one reduction rule for each constructor. The order is -immaterial. For missing constructors, the function is defined to return a -default value. - -If you would like to refer to some rule by name, then you must prefix -the rule with an identifier. These identifiers, like those in the -\texttt{rules} section of a theory, will be visible at the \ML\ level. - -The primitive recursive function can have infix or mixfix syntax: -\begin{ttbox}\underscoreon -consts "@" :: ['a list, 'a list] => 'a list (infixr 60) -primrec - "[] @ ys = ys" - "(x#xs) @ ys = x#(xs @ ys)" -\end{ttbox} - -The reduction rules become part of the default simpset, which -leads to short proof scripts: -\begin{ttbox}\underscoreon -Goal "(xs @ ys) @ zs = xs @ (ys @ zs)"; -by (induct\_tac "xs" 1); -by (ALLGOALS Asm\_simp\_tac); -\end{ttbox} - -\subsubsection{Example: Evaluation of expressions} -Using mutual primitive recursion, we can define evaluation functions \texttt{evala} -and \texttt{eval_bexp} for the datatypes of arithmetic and boolean expressions mentioned in -{\S}\ref{subsec:datatype:basics}: -\begin{ttbox} -consts - evala :: "['a => nat, 'a aexp] => nat" - evalb :: "['a => nat, 'a bexp] => bool" - -primrec - "evala env (If_then_else b a1 a2) = - (if evalb env b then evala env a1 else evala env a2)" - "evala env (Sum a1 a2) = evala env a1 + evala env a2" - "evala env (Diff a1 a2) = evala env a1 - evala env a2" - "evala env (Var v) = env v" - "evala env (Num n) = n" - - "evalb env (Less a1 a2) = (evala env a1 < evala env a2)" - "evalb env (And b1 b2) = (evalb env b1 & evalb env b2)" - "evalb env (Or b1 b2) = (evalb env b1 & evalb env b2)" -\end{ttbox} -Since the value of an expression depends on the value of its variables, -the functions \texttt{evala} and \texttt{evalb} take an additional -parameter, an {\em environment} of type \texttt{'a => nat}, which maps -variables to their values. - -Similarly, we may define substitution functions \texttt{substa} -and \texttt{substb} for expressions: The mapping \texttt{f} of type -\texttt{'a => 'a aexp} given as a parameter is lifted canonically -on the types \texttt{'a aexp} and \texttt{'a bexp}: -\begin{ttbox} -consts - substa :: "['a => 'b aexp, 'a aexp] => 'b aexp" - substb :: "['a => 'b aexp, 'a bexp] => 'b bexp" - -primrec - "substa f (If_then_else b a1 a2) = - If_then_else (substb f b) (substa f a1) (substa f a2)" - "substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)" - "substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)" - "substa f (Var v) = f v" - "substa f (Num n) = Num n" - - "substb f (Less a1 a2) = Less (substa f a1) (substa f a2)" - "substb f (And b1 b2) = And (substb f b1) (substb f b2)" - "substb f (Or b1 b2) = Or (substb f b1) (substb f b2)" -\end{ttbox} -In textbooks about semantics one often finds {\em substitution theorems}, -which express the relationship between substitution and evaluation. For -\texttt{'a aexp} and \texttt{'a bexp}, we can prove such a theorem by mutual -induction, followed by simplification: -\begin{ttbox} -Goal - "evala env (substa (Var(v := a')) a) = - evala (env(v := evala env a')) a & - evalb env (substb (Var(v := a')) b) = - evalb (env(v := evala env a')) b"; -by (induct_tac "a b" 1); -by (ALLGOALS Asm_full_simp_tac); -\end{ttbox} - -\subsubsection{Example: A substitution function for terms} -Functions on datatypes with nested recursion, such as the type -\texttt{term} mentioned in {\S}\ref{subsec:datatype:basics}, are -also defined by mutual primitive recursion. A substitution -function \texttt{subst_term} on type \texttt{term}, similar to the functions -\texttt{substa} and \texttt{substb} described above, can -be defined as follows: -\begin{ttbox} -consts - subst_term :: "['a => ('a,'b) term, ('a,'b) term] => ('a,'b) term" - subst_term_list :: - "['a => ('a,'b) term, ('a,'b) term list] => ('a,'b) term list" - -primrec - "subst_term f (Var a) = f a" - "subst_term f (App b ts) = App b (subst_term_list f ts)" - - "subst_term_list f [] = []" - "subst_term_list f (t # ts) = - subst_term f t # subst_term_list f ts" -\end{ttbox} -The recursion scheme follows the structure of the unfolded definition of type -\texttt{term} shown in {\S}\ref{subsec:datatype:basics}. To prove properties of -this substitution function, mutual induction is needed: -\begin{ttbox} -Goal - "(subst_term ((subst_term f1) o f2) t) = - (subst_term f1 (subst_term f2 t)) & - (subst_term_list ((subst_term f1) o f2) ts) = - (subst_term_list f1 (subst_term_list f2 ts))"; -by (induct_tac "t ts" 1); -by (ALLGOALS Asm_full_simp_tac); -\end{ttbox} - -\subsubsection{Example: A map function for infinitely branching trees} -Defining functions on infinitely branching datatypes by primitive -recursion is just as easy. For example, we can define a function -\texttt{map_tree} on \texttt{'a tree} as follows: -\begin{ttbox} -consts - map_tree :: "('a => 'b) => 'a tree => 'b tree" - -primrec - "map_tree f (Atom a) = Atom (f a)" - "map_tree f (Branch ts) = Branch (\%x. map_tree f (ts x))" -\end{ttbox} -Note that all occurrences of functions such as \texttt{ts} in the -\texttt{primrec} clauses must be applied to an argument. In particular, -\texttt{map_tree f o ts} is not allowed. - -\index{recursion!primitive|)} -\index{*primrec|)} - - -\subsection{General recursive functions} -\label{sec:HOL:recdef} +\section{Old-style recursive function definitions}\label{sec:HOL:recursive} \index{recursion!general|(} \index{*recdef|(} +Old-style recursive definitions via \texttt{recdef} requires that you +supply a well-founded relation that governs the recursion. Recursive +calls are only allowed if they make the argument decrease under the +relation. Complicated recursion forms, such as nested recursion, can +be dealt with. Termination can even be proved at a later time, though +having unsolved termination conditions around can make work +difficult.% +\footnote{This facility is based on Konrad Slind's TFL + package~\cite{slind-tfl}. Thanks are due to Konrad for implementing + TFL and assisting with its installation.} + Using \texttt{recdef}, you can declare functions involving nested recursion and pattern-matching. Recursion need not involve datatypes and there are few syntactic restrictions. Termination is proved by showing that each recursive @@ -2543,186 +1980,6 @@ \index{*recdef|)} -\section{Inductive and coinductive definitions} -\index{*inductive|(} -\index{*coinductive|(} - -An {\bf inductive definition} specifies the least set~$R$ closed under given -rules. (Applying a rule to elements of~$R$ yields a result within~$R$.) For -example, a structural operational semantics is an inductive definition of an -evaluation relation. Dually, a {\bf coinductive definition} specifies the -greatest set~$R$ consistent with given rules. (Every element of~$R$ can be -seen as arising by applying a rule to elements of~$R$.) An important example -is using bisimulation relations to formalise equivalence of processes and -infinite data structures. - -A theory file may contain any number of inductive and coinductive -definitions. They may be intermixed with other declarations; in -particular, the (co)inductive sets {\bf must} be declared separately as -constants, and may have mixfix syntax or be subject to syntax translations. - -Each (co)inductive definition adds definitions to the theory and also -proves some theorems. Each definition creates an \ML\ structure, which is a -substructure of the main theory structure. - -This package is related to the ZF one, described in a separate -paper,% -\footnote{It appeared in CADE~\cite{paulson-CADE}; a longer version is - distributed with Isabelle.} % -which you should refer to in case of difficulties. The package is simpler -than ZF's thanks to HOL's extra-logical automatic type-checking. The types of -the (co)inductive sets determine the domain of the fixedpoint definition, and -the package does not have to use inference rules for type-checking. - - -\subsection{The result structure} -Many of the result structure's components have been discussed in the paper; -others are self-explanatory. -\begin{description} -\item[\tt defs] is the list of definitions of the recursive sets. - -\item[\tt mono] is a monotonicity theorem for the fixedpoint operator. - -\item[\tt unfold] is a fixedpoint equation for the recursive set (the union of -the recursive sets, in the case of mutual recursion). - -\item[\tt intrs] is the list of introduction rules, now proved as theorems, for -the recursive sets. The rules are also available individually, using the -names given them in the theory file. - -\item[\tt elims] is the list of elimination rule. This is for compatibility - with ML scripts; within the theory the name is \texttt{cases}. - -\item[\tt elim] is the head of the list \texttt{elims}. This is for - compatibility only. - -\item[\tt mk_cases] is a function to create simplified instances of {\tt -elim} using freeness reasoning on underlying datatypes. -\end{description} - -For an inductive definition, the result structure contains the -rule \texttt{induct}. For a -coinductive definition, it contains the rule \verb|coinduct|. - -Figure~\ref{def-result-fig} summarises the two result signatures, -specifying the types of all these components. - -\begin{figure} -\begin{ttbox} -sig -val defs : thm list -val mono : thm -val unfold : thm -val intrs : thm list -val elims : thm list -val elim : thm -val mk_cases : string -> thm -{\it(Inductive definitions only)} -val induct : thm -{\it(coinductive definitions only)} -val coinduct : thm -end -\end{ttbox} -\hrule -\caption{The {\ML} result of a (co)inductive definition} \label{def-result-fig} -\end{figure} - -\subsection{The syntax of a (co)inductive definition} -An inductive definition has the form -\begin{ttbox} -inductive {\it inductive sets} - intrs {\it introduction rules} - monos {\it monotonicity theorems} -\end{ttbox} -A coinductive definition is identical, except that it starts with the keyword -\texttt{coinductive}. - -The \texttt{monos} section is optional; if present it is specified by a list -of identifiers. - -\begin{itemize} -\item The \textit{inductive sets} are specified by one or more strings. - -\item The \textit{introduction rules} specify one or more introduction rules in - the form \textit{ident\/}~\textit{string}, where the identifier gives the name of - the rule in the result structure. - -\item The \textit{monotonicity theorems} are required for each operator - applied to a recursive set in the introduction rules. There {\bf must} - be a theorem of the form $A\subseteq B\Imp M(A)\subseteq M(B)$, for each - premise $t\in M(R@i)$ in an introduction rule! - -\item The \textit{constructor definitions} contain definitions of constants - appearing in the introduction rules. In most cases it can be omitted. -\end{itemize} - - -\subsection{*Monotonicity theorems} - -Each theory contains a default set of theorems that are used in monotonicity -proofs. New rules can be added to this set via the \texttt{mono} attribute. -Theory \texttt{Inductive} shows how this is done. In general, the following -monotonicity theorems may be added: -\begin{itemize} -\item Theorems of the form $A\subseteq B\Imp M(A)\subseteq M(B)$, for proving - monotonicity of inductive definitions whose introduction rules have premises - involving terms such as $t\in M(R@i)$. -\item Monotonicity theorems for logical operators, which are of the general form - $\List{\cdots \to \cdots;~\ldots;~\cdots \to \cdots} \Imp - \cdots \to \cdots$. - For example, in the case of the operator $\lor$, the corresponding theorem is - \[ - \infer{P@1 \lor P@2 \to Q@1 \lor Q@2} - {P@1 \to Q@1 & P@2 \to Q@2} - \] -\item De Morgan style equations for reasoning about the ``polarity'' of expressions, e.g. - \[ - (\lnot \lnot P) ~=~ P \qquad\qquad - (\lnot (P \land Q)) ~=~ (\lnot P \lor \lnot Q) - \] -\item Equations for reducing complex operators to more primitive ones whose - monotonicity can easily be proved, e.g. - \[ - (P \to Q) ~=~ (\lnot P \lor Q) \qquad\qquad - \mathtt{Ball}~A~P ~\equiv~ \forall x.~x \in A \to P~x - \] -\end{itemize} - -\subsection{Example of an inductive definition} -Two declarations, included in a theory file, define the finite powerset -operator. First we declare the constant~\texttt{Fin}. Then we declare it -inductively, with two introduction rules: -\begin{ttbox} -consts Fin :: 'a set => 'a set set -inductive "Fin A" - intrs - emptyI "{\ttlbrace}{\ttrbrace} : Fin A" - insertI "[| a: A; b: Fin A |] ==> insert a b : Fin A" -\end{ttbox} -The resulting theory structure contains a substructure, called~\texttt{Fin}. -It contains the \texttt{Fin}$~A$ introduction rules as the list \texttt{Fin.intrs}, -and also individually as \texttt{Fin.emptyI} and \texttt{Fin.consI}. The induction -rule is \texttt{Fin.induct}. - -For another example, here is a theory file defining the accessible part of a -relation. The paper \cite{paulson-CADE} discusses a ZF version of this -example in more detail. -\begin{ttbox} -Acc = WF + Inductive + - -consts acc :: "('a * 'a)set => 'a set" (* accessible part *) - -inductive "acc r" - intrs - accI "ALL y. (y, x) : r --> y : acc r ==> x : acc r" - -end -\end{ttbox} -The Isabelle distribution contains many other inductive definitions. - -\index{*coinductive|)} \index{*inductive|)} - - \section{Example: Cantor's Theorem}\label{sec:hol-cantor} Cantor's Theorem states that every set has more subsets than it has elements. It has become a favourite example in higher-order logic since diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarImplementation/Thy/Logic.thy --- a/doc-src/IsarImplementation/Thy/Logic.thy Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarImplementation/Thy/Logic.thy Wed Jun 08 10:24:07 2011 +0200 @@ -366,6 +366,7 @@ @{index_ML fastype_of: "term -> typ"} \\ @{index_ML lambda: "term -> term -> term"} \\ @{index_ML betapply: "term * term -> term"} \\ + @{index_ML incr_boundvars: "int -> term -> term"} \\ @{index_ML Sign.declare_const: "Proof.context -> (binding * typ) * mixfix -> theory -> term * theory"} \\ @{index_ML Sign.add_abbrev: "string -> binding * term -> @@ -414,6 +415,12 @@ "t u"}, with topmost @{text "\"}-conversion if @{text "t"} is an abstraction. + \item @{ML incr_boundvars}~@{text "j"} increments a term's dangling + bound variables by the offset @{text "j"}. This is required when + moving a subterm into a context where it is enclosed by a different + number of abstractions. Bound variables with a matching abstraction + are unaffected. + \item @{ML Sign.declare_const}~@{text "ctxt ((c, \), mx)"} declares a new constant @{text "c :: \"} with optional mixfix syntax. @@ -634,6 +641,7 @@ \begin{mldecls} @{index_ML_type thm} \\ @{index_ML proofs: "int Unsynchronized.ref"} \\ + @{index_ML Thm.transfer: "theory -> thm -> thm"} \\ @{index_ML Thm.assume: "cterm -> thm"} \\ @{index_ML Thm.forall_intr: "cterm -> thm -> thm"} \\ @{index_ML Thm.forall_elim: "cterm -> thm -> thm"} \\ @@ -682,6 +690,13 @@ records full proof terms. Officially named theorems that contribute to a result are recorded in any case. + \item @{ML Thm.transfer}~@{text "thy thm"} transfers the given + theorem to a \emph{larger} theory, see also \secref{sec:context}. + This formal adjustment of the background context has no logical + significance, but is occasionally required for formal reasons, e.g.\ + when theorems that are imported from more basic theories are used in + the current situation. + \item @{ML Thm.assume}, @{ML Thm.forall_intr}, @{ML Thm.forall_elim}, @{ML Thm.implies_intr}, and @{ML Thm.implies_elim} correspond to the primitive inferences of \figref{fig:prim-rules}. diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarImplementation/Thy/document/Logic.tex --- a/doc-src/IsarImplementation/Thy/document/Logic.tex Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarImplementation/Thy/document/Logic.tex Wed Jun 08 10:24:07 2011 +0200 @@ -404,6 +404,7 @@ \indexdef{}{ML}{fastype\_of}\verb|fastype_of: term -> typ| \\ \indexdef{}{ML}{lambda}\verb|lambda: term -> term -> term| \\ \indexdef{}{ML}{betapply}\verb|betapply: term * term -> term| \\ + \indexdef{}{ML}{incr\_boundvars}\verb|incr_boundvars: int -> term -> term| \\ \indexdef{}{ML}{Sign.declare\_const}\verb|Sign.declare_const: Proof.context ->|\isasep\isanewline% \verb| (binding * typ) * mixfix -> theory -> term * theory| \\ \indexdef{}{ML}{Sign.add\_abbrev}\verb|Sign.add_abbrev: string -> binding * term ->|\isasep\isanewline% @@ -444,6 +445,12 @@ \item \verb|betapply|~\isa{{\isaliteral{28}{\isacharparenleft}}t{\isaliteral{2C}{\isacharcomma}}\ u{\isaliteral{29}{\isacharparenright}}} produces an application \isa{t\ u}, with topmost \isa{{\isaliteral{5C3C626574613E}{\isasymbeta}}}-conversion if \isa{t} is an abstraction. + \item \verb|incr_boundvars|~\isa{j} increments a term's dangling + bound variables by the offset \isa{j}. This is required when + moving a subterm into a context where it is enclosed by a different + number of abstractions. Bound variables with a matching abstraction + are unaffected. + \item \verb|Sign.declare_const|~\isa{ctxt\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}c{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ mx{\isaliteral{29}{\isacharparenright}}} declares a new constant \isa{c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}} with optional mixfix syntax. @@ -705,6 +712,7 @@ \begin{mldecls} \indexdef{}{ML type}{thm}\verb|type thm| \\ \indexdef{}{ML}{proofs}\verb|proofs: int Unsynchronized.ref| \\ + \indexdef{}{ML}{Thm.transfer}\verb|Thm.transfer: theory -> thm -> thm| \\ \indexdef{}{ML}{Thm.assume}\verb|Thm.assume: cterm -> thm| \\ \indexdef{}{ML}{Thm.forall\_intr}\verb|Thm.forall_intr: cterm -> thm -> thm| \\ \indexdef{}{ML}{Thm.forall\_elim}\verb|Thm.forall_elim: cterm -> thm -> thm| \\ @@ -752,6 +760,13 @@ records full proof terms. Officially named theorems that contribute to a result are recorded in any case. + \item \verb|Thm.transfer|~\isa{thy\ thm} transfers the given + theorem to a \emph{larger} theory, see also \secref{sec:context}. + This formal adjustment of the background context has no logical + significance, but is occasionally required for formal reasons, e.g.\ + when theorems that are imported from more basic theories are used in + the current situation. + \item \verb|Thm.assume|, \verb|Thm.forall_intr|, \verb|Thm.forall_elim|, \verb|Thm.implies_intr|, and \verb|Thm.implies_elim| correspond to the primitive inferences of \figref{fig:prim-rules}. diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/IsaMakefile --- a/doc-src/IsarRef/IsaMakefile Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/IsaMakefile Wed Jun 08 10:24:07 2011 +0200 @@ -23,10 +23,10 @@ $(LOG)/HOL-IsarRef.gz: Thy/ROOT.ML ../antiquote_setup.ML Thy/Base.thy \ Thy/First_Order_Logic.thy Thy/Framework.thy Thy/Inner_Syntax.thy \ - Thy/Introduction.thy Thy/Outer_Syntax.thy Thy/Spec.thy Thy/Proof.thy \ + Thy/Preface.thy Thy/Outer_Syntax.thy Thy/Spec.thy Thy/Proof.thy \ Thy/Misc.thy Thy/Document_Preparation.thy Thy/Generic.thy \ - Thy/HOL_Specific.thy Thy/Quick_Reference.thy Thy/Symbols.thy \ - Thy/ML_Tactic.thy + Thy/HOL_Specific.thy Thy/Quick_Reference.thy Thy/Synopsis.thy \ + Thy/Symbols.thy Thy/ML_Tactic.thy @$(USEDIR) -s IsarRef HOL Thy @rm -f Thy/document/isabelle.sty Thy/document/isabellesym.sty \ Thy/document/pdfsetup.sty Thy/document/session.tex diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Makefile --- a/doc-src/IsarRef/Makefile Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/Makefile Wed Jun 08 10:24:07 2011 +0200 @@ -9,12 +9,13 @@ NAME = isar-ref -FILES = isar-ref.tex style.sty Thy/document/Generic.tex \ - Thy/document/HOLCF_Specific.tex Thy/document/HOL_Specific.tex \ - Thy/document/ML_Tactic.tex Thy/document/Proof.tex \ - Thy/document/Quick_Reference.tex Thy/document/Spec.tex \ +FILES = isar-ref.tex style.sty Thy/document/Framework.tex \ + Thy/document/Generic.tex Thy/document/HOLCF_Specific.tex \ + Thy/document/HOL_Specific.tex Thy/document/ML_Tactic.tex \ + Thy/document/Proof.tex Thy/document/Quick_Reference.tex \ + Thy/document/Spec.tex Thy/document/Synopsis.tex \ Thy/document/ZF_Specific.tex Thy/document/Inner_Syntax.tex \ - Thy/document/Introduction.tex Thy/document/Document_Preparation.tex \ + Thy/document/Preface.tex Thy/document/Document_Preparation.tex \ Thy/document/Misc.tex Thy/document/Outer_Syntax.tex \ Thy/document/Symbols.tex ../isar.sty ../proof.sty ../iman.sty \ ../extra.sty ../ttbox.sty ../../lib/texinputs/isabelle.sty \ diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/Document_Preparation.thy --- a/doc-src/IsarRef/Thy/Document_Preparation.thy Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/Thy/Document_Preparation.thy Wed Jun 08 10:24:07 2011 +0200 @@ -295,7 +295,7 @@ *} -subsubsection {* Styled antiquotations *} +subsection {* Styled antiquotations *} text {* The antiquotations @{text thm}, @{text prop} and @{text term} admit an extra \emph{style} specification to modify the @@ -323,7 +323,7 @@ *} -subsubsection {* General options *} +subsection {* General options *} text {* The following options are available to tune the printed output of antiquotations. Note that many of these coincide with global ML diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/Generic.thy --- a/doc-src/IsarRef/Thy/Generic.thy Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/Thy/Generic.thy Wed Jun 08 10:24:07 2011 +0200 @@ -468,7 +468,22 @@ subsection {* Simplification procedures *} -text {* +text {* Simplification procedures are ML functions that produce proven + rewrite rules on demand. They are associated with higher-order + patterns that approximate the left-hand sides of equations. The + Simplifier first matches the current redex against one of the LHS + patterns; if this succeeds, the corresponding ML function is + invoked, passing the Simplifier context and redex term. Thus rules + may be specifically fashioned for particular situations, resulting + in a more powerful mechanism than term rewriting by a fixed set of + rules. + + Any successful result needs to be a (possibly conditional) rewrite + rule @{text "t \ u"} that is applicable to the current redex. The + rule will be applied just as any ordinary rewrite rule. It is + expected to be already in \emph{internal form}, bypassing the + automatic preprocessing of object-level equivalences. + \begin{matharray}{rcl} @{command_def "simproc_setup"} & : & @{text "local_theory \ local_theory"} \\ simproc & : & @{text attribute} \\ @@ -512,6 +527,26 @@ *} +subsubsection {* Example *} + +text {* The following simplification procedure for @{thm + [source=false, show_types] unit_eq} in HOL performs fine-grained + control over rule application, beyond higher-order pattern matching. + Declaring @{thm unit_eq} as @{attribute simp} directly would make + the simplifier loop! Note that a version of this simplification + procedure is already active in Isabelle/HOL. *} + +simproc_setup unit ("x::unit") = {* + fn _ => fn _ => fn ct => + if HOLogic.is_unit (term_of ct) then NONE + else SOME (mk_meta_eq @{thm unit_eq}) +*} + +text {* Since the Simplifier applies simplification procedures + frequently, it is important to make the failure check in ML + reasonably fast. *} + + subsection {* Forward simplification *} text {* @@ -547,7 +582,462 @@ section {* The Classical Reasoner \label{sec:classical} *} -subsection {* Basic methods *} +subsection {* Basic concepts *} + +text {* Although Isabelle is generic, many users will be working in + some extension of classical first-order logic. Isabelle/ZF is built + upon theory FOL, while Isabelle/HOL conceptually contains + first-order logic as a fragment. Theorem-proving in predicate logic + is undecidable, but many automated strategies have been developed to + assist in this task. + + Isabelle's classical reasoner is a generic package that accepts + certain information about a logic and delivers a suite of automatic + proof tools, based on rules that are classified and declared in the + context. These proof procedures are slow and simplistic compared + with high-end automated theorem provers, but they can save + considerable time and effort in practice. They can prove theorems + such as Pelletier's \cite{pelletier86} problems 40 and 41 in a few + milliseconds (including full proof reconstruction): *} + +lemma "(\y. \x. F x y \ F x x) \ \ (\x. \y. \z. F z y \ \ F z x)" + by blast + +lemma "(\z. \y. \x. f x y \ f x z \ \ f x x) \ \ (\z. \x. f x z)" + by blast + +text {* The proof tools are generic. They are not restricted to + first-order logic, and have been heavily used in the development of + the Isabelle/HOL library and applications. The tactics can be + traced, and their components can be called directly; in this manner, + any proof can be viewed interactively. *} + + +subsubsection {* The sequent calculus *} + +text {* Isabelle supports natural deduction, which is easy to use for + interactive proof. But natural deduction does not easily lend + itself to automation, and has a bias towards intuitionism. For + certain proofs in classical logic, it can not be called natural. + The \emph{sequent calculus}, a generalization of natural deduction, + is easier to automate. + + A \textbf{sequent} has the form @{text "\ \ \"}, where @{text "\"} + and @{text "\"} are sets of formulae.\footnote{For first-order + logic, sequents can equivalently be made from lists or multisets of + formulae.} The sequent @{text "P\<^sub>1, \, P\<^sub>m \ Q\<^sub>1, \, Q\<^sub>n"} is + \textbf{valid} if @{text "P\<^sub>1 \ \ \ P\<^sub>m"} implies @{text "Q\<^sub>1 \ \ \ + Q\<^sub>n"}. Thus @{text "P\<^sub>1, \, P\<^sub>m"} represent assumptions, each of which + is true, while @{text "Q\<^sub>1, \, Q\<^sub>n"} represent alternative goals. A + sequent is \textbf{basic} if its left and right sides have a common + formula, as in @{text "P, Q \ Q, R"}; basic sequents are trivially + valid. + + Sequent rules are classified as \textbf{right} or \textbf{left}, + indicating which side of the @{text "\"} symbol they operate on. + Rules that operate on the right side are analogous to natural + deduction's introduction rules, and left rules are analogous to + elimination rules. The sequent calculus analogue of @{text "(\I)"} + is the rule + \[ + \infer[@{text "(\R)"}]{@{text "\ \ \, P \ Q"}}{@{text "P, \ \ \, Q"}} + \] + Applying the rule backwards, this breaks down some implication on + the right side of a sequent; @{text "\"} and @{text "\"} stand for + the sets of formulae that are unaffected by the inference. The + analogue of the pair @{text "(\I1)"} and @{text "(\I2)"} is the + single rule + \[ + \infer[@{text "(\R)"}]{@{text "\ \ \, P \ Q"}}{@{text "\ \ \, P, Q"}} + \] + This breaks down some disjunction on the right side, replacing it by + both disjuncts. Thus, the sequent calculus is a kind of + multiple-conclusion logic. + + To illustrate the use of multiple formulae on the right, let us + prove the classical theorem @{text "(P \ Q) \ (Q \ P)"}. Working + backwards, we reduce this formula to a basic sequent: + \[ + \infer[@{text "(\R)"}]{@{text "\ (P \ Q) \ (Q \ P)"}} + {\infer[@{text "(\R)"}]{@{text "\ (P \ Q), (Q \ P)"}} + {\infer[@{text "(\R)"}]{@{text "P \ Q, (Q \ P)"}} + {@{text "P, Q \ Q, P"}}}} + \] + + This example is typical of the sequent calculus: start with the + desired theorem and apply rules backwards in a fairly arbitrary + manner. This yields a surprisingly effective proof procedure. + Quantifiers add only few complications, since Isabelle handles + parameters and schematic variables. See \cite[Chapter + 10]{paulson-ml2} for further discussion. *} + + +subsubsection {* Simulating sequents by natural deduction *} + +text {* Isabelle can represent sequents directly, as in the + object-logic LK. But natural deduction is easier to work with, and + most object-logics employ it. Fortunately, we can simulate the + sequent @{text "P\<^sub>1, \, P\<^sub>m \ Q\<^sub>1, \, Q\<^sub>n"} by the Isabelle formula + @{text "P\<^sub>1 \ \ \ P\<^sub>m \ \ Q\<^sub>2 \ ... \ \ Q\<^sub>n \ Q\<^sub>1"} where the order of + the assumptions and the choice of @{text "Q\<^sub>1"} are arbitrary. + Elim-resolution plays a key role in simulating sequent proofs. + + We can easily handle reasoning on the left. Elim-resolution with + the rules @{text "(\E)"}, @{text "(\E)"} and @{text "(\E)"} achieves + a similar effect as the corresponding sequent rules. For the other + connectives, we use sequent-style elimination rules instead of + destruction rules such as @{text "(\E1, 2)"} and @{text "(\E)"}. + But note that the rule @{text "(\L)"} has no effect under our + representation of sequents! + \[ + \infer[@{text "(\L)"}]{@{text "\ P, \ \ \"}}{@{text "\ \ \, P"}} + \] + + What about reasoning on the right? Introduction rules can only + affect the formula in the conclusion, namely @{text "Q\<^sub>1"}. The + other right-side formulae are represented as negated assumptions, + @{text "\ Q\<^sub>2, \, \ Q\<^sub>n"}. In order to operate on one of these, it + must first be exchanged with @{text "Q\<^sub>1"}. Elim-resolution with the + @{text swap} rule has this effect: @{text "\ P \ (\ R \ P) \ R"} + + To ensure that swaps occur only when necessary, each introduction + rule is converted into a swapped form: it is resolved with the + second premise of @{text "(swap)"}. The swapped form of @{text + "(\I)"}, which might be called @{text "(\\E)"}, is + @{text [display] "\ (P \ Q) \ (\ R \ P) \ (\ R \ Q) \ R"} + + Similarly, the swapped form of @{text "(\I)"} is + @{text [display] "\ (P \ Q) \ (\ R \ P \ Q) \ R"} + + Swapped introduction rules are applied using elim-resolution, which + deletes the negated formula. Our representation of sequents also + requires the use of ordinary introduction rules. If we had no + regard for readability of intermediate goal states, we could treat + the right side more uniformly by representing sequents as @{text + [display] "P\<^sub>1 \ \ \ P\<^sub>m \ \ Q\<^sub>1 \ \ \ \ Q\<^sub>n \ \"} +*} + + +subsubsection {* Extra rules for the sequent calculus *} + +text {* As mentioned, destruction rules such as @{text "(\E1, 2)"} and + @{text "(\E)"} must be replaced by sequent-style elimination rules. + In addition, we need rules to embody the classical equivalence + between @{text "P \ Q"} and @{text "\ P \ Q"}. The introduction + rules @{text "(\I1, 2)"} are replaced by a rule that simulates + @{text "(\R)"}: @{text [display] "(\ Q \ P) \ P \ Q"} + + The destruction rule @{text "(\E)"} is replaced by @{text [display] + "(P \ Q) \ (\ P \ R) \ (Q \ R) \ R"} + + Quantifier replication also requires special rules. In classical + logic, @{text "\x. P x"} is equivalent to @{text "\ (\x. \ P x)"}; + the rules @{text "(\R)"} and @{text "(\L)"} are dual: + \[ + \infer[@{text "(\R)"}]{@{text "\ \ \, \x. P x"}}{@{text "\ \ \, \x. P x, P t"}} + \qquad + \infer[@{text "(\L)"}]{@{text "\x. P x, \ \ \"}}{@{text "P t, \x. P x, \ \ \"}} + \] + Thus both kinds of quantifier may be replicated. Theorems requiring + multiple uses of a universal formula are easy to invent; consider + @{text [display] "(\x. P x \ P (f x)) \ P a \ P (f\<^sup>n a)"} for any + @{text "n > 1"}. Natural examples of the multiple use of an + existential formula are rare; a standard one is @{text "\x. \y. P x + \ P y"}. + + Forgoing quantifier replication loses completeness, but gains + decidability, since the search space becomes finite. Many useful + theorems can be proved without replication, and the search generally + delivers its verdict in a reasonable time. To adopt this approach, + represent the sequent rules @{text "(\R)"}, @{text "(\L)"} and + @{text "(\R)"} by @{text "(\I)"}, @{text "(\E)"} and @{text "(\I)"}, + respectively, and put @{text "(\E)"} into elimination form: @{text + [display] "\x. P x \ (P t \ Q) \ Q"} + + Elim-resolution with this rule will delete the universal formula + after a single use. To replicate universal quantifiers, replace the + rule by @{text [display] "\x. P x \ (P t \ \x. P x \ Q) \ Q"} + + To replicate existential quantifiers, replace @{text "(\I)"} by + @{text [display] "(\ (\x. P x) \ P t) \ \x. P x"} + + All introduction rules mentioned above are also useful in swapped + form. + + Replication makes the search space infinite; we must apply the rules + with care. The classical reasoner distinguishes between safe and + unsafe rules, applying the latter only when there is no alternative. + Depth-first search may well go down a blind alley; best-first search + is better behaved in an infinite search space. However, quantifier + replication is too expensive to prove any but the simplest theorems. +*} + + +subsection {* Rule declarations *} + +text {* The proof tools of the Classical Reasoner depend on + collections of rules declared in the context, which are classified + as introduction, elimination or destruction and as \emph{safe} or + \emph{unsafe}. In general, safe rules can be attempted blindly, + while unsafe rules must be used with care. A safe rule must never + reduce a provable goal to an unprovable set of subgoals. + + The rule @{text "P \ P \ Q"} is unsafe because it reduces @{text "P + \ Q"} to @{text "P"}, which might turn out as premature choice of an + unprovable subgoal. Any rule is unsafe whose premises contain new + unknowns. The elimination rule @{text "\x. P x \ (P t \ Q) \ Q"} is + unsafe, since it is applied via elim-resolution, which discards the + assumption @{text "\x. P x"} and replaces it by the weaker + assumption @{text "P t"}. The rule @{text "P t \ \x. P x"} is + unsafe for similar reasons. The quantifier duplication rule @{text + "\x. P x \ (P t \ \x. P x \ Q) \ Q"} is unsafe in a different sense: + since it keeps the assumption @{text "\x. P x"}, it is prone to + looping. In classical first-order logic, all rules are safe except + those mentioned above. + + The safe~/ unsafe distinction is vague, and may be regarded merely + as a way of giving some rules priority over others. One could argue + that @{text "(\E)"} is unsafe, because repeated application of it + could generate exponentially many subgoals. Induction rules are + unsafe because inductive proofs are difficult to set up + automatically. Any inference is unsafe that instantiates an unknown + in the proof state --- thus matching must be used, rather than + unification. Even proof by assumption is unsafe if it instantiates + unknowns shared with other subgoals. + + \begin{matharray}{rcl} + @{command_def "print_claset"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{attribute_def intro} & : & @{text attribute} \\ + @{attribute_def elim} & : & @{text attribute} \\ + @{attribute_def dest} & : & @{text attribute} \\ + @{attribute_def rule} & : & @{text attribute} \\ + @{attribute_def iff} & : & @{text attribute} \\ + @{attribute_def swapped} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}? + ; + @@{attribute rule} 'del' + ; + @@{attribute iff} (((() | 'add') '?'?) | 'del') + "} + + \begin{description} + + \item @{command "print_claset"} prints the collection of rules + declared to the Classical Reasoner, i.e.\ the @{ML_type claset} + within the context. + + \item @{attribute intro}, @{attribute elim}, and @{attribute dest} + declare introduction, elimination, and destruction rules, + respectively. By default, rules are considered as \emph{unsafe} + (i.e.\ not applied blindly without backtracking), while ``@{text + "!"}'' classifies as \emph{safe}. Rule declarations marked by + ``@{text "?"}'' coincide with those of Isabelle/Pure, cf.\ + \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps + of the @{method rule} method). The optional natural number + specifies an explicit weight argument, which is ignored by the + automated reasoning tools, but determines the search order of single + rule steps. + + Introduction rules are those that can be applied using ordinary + resolution. Their swapped forms are generated internally, which + will be applied using elim-resolution. Elimination rules are + applied using elim-resolution. Rules are sorted by the number of + new subgoals they will yield; rules that generate the fewest + subgoals will be tried first. Otherwise, later declarations take + precedence over earlier ones. + + Rules already present in the context with the same classification + are ignored. A warning is printed if the rule has already been + added with some other classification, but the rule is added anyway + as requested. + + \item @{attribute rule}~@{text del} deletes all occurrences of a + rule from the classical context, regardless of its classification as + introduction~/ elimination~/ destruction and safe~/ unsafe. + + \item @{attribute iff} declares logical equivalences to the + Simplifier and the Classical reasoner at the same time. + Non-conditional rules result in a safe introduction and elimination + pair; conditional ones are considered unsafe. Rules with negative + conclusion are automatically inverted (using @{text "\"}-elimination + internally). + + The ``@{text "?"}'' version of @{attribute iff} declares rules to + the Isabelle/Pure context only, and omits the Simplifier + declaration. + + \item @{attribute swapped} turns an introduction rule into an + elimination, by resolving with the classical swap principle @{text + "\ P \ (\ R \ P) \ R"} in the second position. This is mainly for + illustrative purposes: the Classical Reasoner already swaps rules + internally as explained above. + + \end{description} +*} + + +subsection {* Automated methods *} + +text {* + \begin{matharray}{rcl} + @{method_def blast} & : & @{text method} \\ + @{method_def auto} & : & @{text method} \\ + @{method_def force} & : & @{text method} \\ + @{method_def fast} & : & @{text method} \\ + @{method_def slow} & : & @{text method} \\ + @{method_def best} & : & @{text method} \\ + @{method_def fastsimp} & : & @{text method} \\ + @{method_def slowsimp} & : & @{text method} \\ + @{method_def bestsimp} & : & @{text method} \\ + \end{matharray} + + @{rail " + @@{method blast} @{syntax nat}? (@{syntax clamod} * ) + ; + @@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * ) + ; + @@{method force} (@{syntax clasimpmod} * ) + ; + (@@{method fast} | @@{method slow} | @@{method best}) (@{syntax clamod} * ) + ; + (@@{method fastsimp} | @@{method slowsimp} | @@{method bestsimp}) + (@{syntax clasimpmod} * ) + ; + @{syntax_def clamod}: + (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thmrefs} + ; + @{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') | + ('cong' | 'split') (() | 'add' | 'del') | + 'iff' (((() | 'add') '?'?) | 'del') | + (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thmrefs} + "} + + \begin{description} + + \item @{method blast} is a separate classical tableau prover that + uses the same classical rule declarations as explained before. + + Proof search is coded directly in ML using special data structures. + A successful proof is then reconstructed using regular Isabelle + inferences. It is faster and more powerful than the other classical + reasoning tools, but has major limitations too. + + \begin{itemize} + + \item It does not use the classical wrapper tacticals, such as the + integration with the Simplifier of @{method fastsimp}. + + \item It does not perform higher-order unification, as needed by the + rule @{thm [source=false] rangeI} in HOL. There are often + alternatives to such rules, for example @{thm [source=false] + range_eqI}. + + \item Function variables may only be applied to parameters of the + subgoal. (This restriction arises because the prover does not use + higher-order unification.) If other function variables are present + then the prover will fail with the message \texttt{Function Var's + argument not a bound variable}. + + \item Its proof strategy is more general than @{method fast} but can + be slower. If @{method blast} fails or seems to be running forever, + try @{method fast} and the other proof tools described below. + + \end{itemize} + + The optional integer argument specifies a bound for the number of + unsafe steps used in a proof. By default, @{method blast} starts + with a bound of 0 and increases it successively to 20. In contrast, + @{text "(blast lim)"} tries to prove the goal using a search bound + of @{text "lim"}. Sometimes a slow proof using @{method blast} can + be made much faster by supplying the successful search bound to this + proof method instead. + + \item @{method auto} combines classical reasoning with + simplification. It is intended for situations where there are a lot + of mostly trivial subgoals; it proves all the easy ones, leaving the + ones it cannot prove. Occasionally, attempting to prove the hard + ones may take a long time. + + %FIXME auto nat arguments + + \item @{method force} is intended to prove the first subgoal + completely, using many fancy proof tools and performing a rather + exhaustive search. As a result, proof attempts may take rather long + or diverge easily. + + \item @{method fast}, @{method best}, @{method slow} attempt to + prove the first subgoal using sequent-style reasoning as explained + before. Unlike @{method blast}, they construct proofs directly in + Isabelle. + + There is a difference in search strategy and back-tracking: @{method + fast} uses depth-first search and @{method best} uses best-first + search (guided by a heuristic function: normally the total size of + the proof state). + + Method @{method slow} is like @{method fast}, but conducts a broader + search: it may, when backtracking from a failed proof attempt, undo + even the step of proving a subgoal by assumption. + + \item @{method fastsimp}, @{method slowsimp}, @{method bestsimp} are + like @{method fast}, @{method slow}, @{method best}, respectively, + but use the Simplifier as additional wrapper. + + \end{description} + + Any of the above methods support additional modifiers of the context + of classical (and simplifier) rules, but the ones related to the + Simplifier are explicitly prefixed by @{text simp} here. The + semantics of these ad-hoc rule declarations is analogous to the + attributes given before. Facts provided by forward chaining are + inserted into the goal before commencing proof search. +*} + + +subsection {* Semi-automated methods *} + +text {* These proof methods may help in situations when the + fully-automated tools fail. The result is a simpler subgoal that + can be tackled by other means, such as by manual instantiation of + quantifiers. + + \begin{matharray}{rcl} + @{method_def safe} & : & @{text method} \\ + @{method_def clarify} & : & @{text method} \\ + @{method_def clarsimp} & : & @{text method} \\ + \end{matharray} + + @{rail " + (@@{method safe} | @@{method clarify}) (@{syntax clamod} * ) + ; + @@{method clarsimp} (@{syntax clasimpmod} * ) + "} + + \begin{description} + + \item @{method safe} repeatedly performs safe steps on all subgoals. + It is deterministic, with at most one outcome. + + \item @{method clarify} performs a series of safe steps as follows. + + No splitting step is applied; for example, the subgoal @{text "A \ + B"} is left as a conjunction. Proof by assumption, Modus Ponens, + etc., may be performed provided they do not instantiate unknowns. + Assumptions of the form @{text "x = t"} may be eliminated. The safe + wrapper tactical is applied. + + \item @{method clarsimp} acts like @{method clarify}, but also does + simplification. Note that if the Simplifier context includes a + splitter for the premises, the subgoal may still be split. + + \end{description} +*} + + +subsection {* Structured proof methods *} text {* \begin{matharray}{rcl} @@ -564,10 +1054,9 @@ \begin{description} \item @{method rule} as offered by the Classical Reasoner is a - refinement over the primitive one (see \secref{sec:pure-meth-att}). - Both versions essentially work the same, but the classical version - observes the classical rule context in addition to that of - Isabelle/Pure. + refinement over the Pure one (see \secref{sec:pure-meth-att}). Both + versions work the same, but the classical version observes the + classical rule context in addition to that of Isabelle/Pure. Common object logics (HOL, ZF, etc.) declare a rich collection of classical rules (even if these would qualify as intuitionistic @@ -589,167 +1078,6 @@ *} -subsection {* Automated methods *} - -text {* - \begin{matharray}{rcl} - @{method_def blast} & : & @{text method} \\ - @{method_def fast} & : & @{text method} \\ - @{method_def slow} & : & @{text method} \\ - @{method_def best} & : & @{text method} \\ - @{method_def safe} & : & @{text method} \\ - @{method_def clarify} & : & @{text method} \\ - \end{matharray} - - @{rail " - @@{method blast} @{syntax nat}? (@{syntax clamod} * ) - ; - (@@{method fast} | @@{method slow} | @@{method best} | @@{method safe} | @@{method clarify}) - (@{syntax clamod} * ) - ; - - @{syntax_def clamod}: - (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thmrefs} - "} - - \begin{description} - - \item @{method blast} refers to the classical tableau prover (see - @{ML blast_tac} in \cite{isabelle-ref}). The optional argument - specifies a user-supplied search bound (default 20). - - \item @{method fast}, @{method slow}, @{method best}, @{method - safe}, and @{method clarify} refer to the generic classical - reasoner. See @{ML fast_tac}, @{ML slow_tac}, @{ML best_tac}, @{ML - safe_tac}, and @{ML clarify_tac} in \cite{isabelle-ref} for more - information. - - \end{description} - - Any of the above methods support additional modifiers of the context - of classical rules. Their semantics is analogous to the attributes - given before. Facts provided by forward chaining are inserted into - the goal before commencing proof search. -*} - - -subsection {* Combined automated methods \label{sec:clasimp} *} - -text {* - \begin{matharray}{rcl} - @{method_def auto} & : & @{text method} \\ - @{method_def force} & : & @{text method} \\ - @{method_def clarsimp} & : & @{text method} \\ - @{method_def fastsimp} & : & @{text method} \\ - @{method_def slowsimp} & : & @{text method} \\ - @{method_def bestsimp} & : & @{text method} \\ - \end{matharray} - - @{rail " - @@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * ) - ; - (@@{method force} | @@{method clarsimp} | @@{method fastsimp} | @@{method slowsimp} | - @@{method bestsimp}) (@{syntax clasimpmod} * ) - ; - - @{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') | - ('cong' | 'split') (() | 'add' | 'del') | - 'iff' (((() | 'add') '?'?) | 'del') | - (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thmrefs} - "} - - \begin{description} - - \item @{method auto}, @{method force}, @{method clarsimp}, @{method - fastsimp}, @{method slowsimp}, and @{method bestsimp} provide access - to Isabelle's combined simplification and classical reasoning - tactics. These correspond to @{ML auto_tac}, @{ML force_tac}, @{ML - clarsimp_tac}, and Classical Reasoner tactics with the Simplifier - added as wrapper, see \cite{isabelle-ref} for more information. The - modifier arguments correspond to those given in - \secref{sec:simplifier} and \secref{sec:classical}. Just note that - the ones related to the Simplifier are prefixed by @{text simp} - here. - - Facts provided by forward chaining are inserted into the goal before - doing the search. - - \end{description} -*} - - -subsection {* Declaring rules *} - -text {* - \begin{matharray}{rcl} - @{command_def "print_claset"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{attribute_def intro} & : & @{text attribute} \\ - @{attribute_def elim} & : & @{text attribute} \\ - @{attribute_def dest} & : & @{text attribute} \\ - @{attribute_def rule} & : & @{text attribute} \\ - @{attribute_def iff} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}? - ; - @@{attribute rule} 'del' - ; - @@{attribute iff} (((() | 'add') '?'?) | 'del') - "} - - \begin{description} - - \item @{command "print_claset"} prints the collection of rules - declared to the Classical Reasoner, which is also known as - ``claset'' internally \cite{isabelle-ref}. - - \item @{attribute intro}, @{attribute elim}, and @{attribute dest} - declare introduction, elimination, and destruction rules, - respectively. By default, rules are considered as \emph{unsafe} - (i.e.\ not applied blindly without backtracking), while ``@{text - "!"}'' classifies as \emph{safe}. Rule declarations marked by - ``@{text "?"}'' coincide with those of Isabelle/Pure, cf.\ - \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps - of the @{method rule} method). The optional natural number - specifies an explicit weight argument, which is ignored by automated - tools, but determines the search order of single rule steps. - - \item @{attribute rule}~@{text del} deletes introduction, - elimination, or destruction rules from the context. - - \item @{attribute iff} declares logical equivalences to the - Simplifier and the Classical reasoner at the same time. - Non-conditional rules result in a ``safe'' introduction and - elimination pair; conditional ones are considered ``unsafe''. Rules - with negative conclusion are automatically inverted (using @{text - "\"}-elimination internally). - - The ``@{text "?"}'' version of @{attribute iff} declares rules to - the Isabelle/Pure context only, and omits the Simplifier - declaration. - - \end{description} -*} - - -subsection {* Classical operations *} - -text {* - \begin{matharray}{rcl} - @{attribute_def swapped} & : & @{text attribute} \\ - \end{matharray} - - \begin{description} - - \item @{attribute swapped} turns an introduction rule into an - elimination, by resolving with the classical swap principle @{text - "(\ B \ A) \ (\ A \ B)"}. - - \end{description} -*} - - section {* Object-logic setup \label{sec:object-logic} *} text {* diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/HOL_Specific.thy --- a/doc-src/IsarRef/Thy/HOL_Specific.thy Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/Thy/HOL_Specific.thy Wed Jun 08 10:24:07 2011 +0200 @@ -4,88 +4,764 @@ chapter {* Isabelle/HOL \label{ch:hol} *} -section {* Typedef axiomatization \label{sec:hol-typedef} *} +section {* Higher-Order Logic *} + +text {* Isabelle/HOL is based on Higher-Order Logic, a polymorphic + version of Church's Simple Theory of Types. HOL can be best + understood as a simply-typed version of classical set theory. The + logic was first implemented in Gordon's HOL system + \cite{mgordon-hol}. It extends Church's original logic + \cite{church40} by explicit type variables (naive polymorphism) and + a sound axiomatization scheme for new types based on subsets of + existing types. + + Andrews's book \cite{andrews86} is a full description of the + original Church-style higher-order logic, with proofs of correctness + and completeness wrt.\ certain set-theoretic interpretations. The + particular extensions of Gordon-style HOL are explained semantically + in two chapters of the 1993 HOL book \cite{pitts93}. + + Experience with HOL over decades has demonstrated that higher-order + logic is widely applicable in many areas of mathematics and computer + science. In a sense, Higher-Order Logic is simpler than First-Order + Logic, because there are fewer restrictions and special cases. Note + that HOL is \emph{weaker} than FOL with axioms for ZF set theory, + which is traditionally considered the standard foundation of regular + mathematics, but for most applications this does not matter. If you + prefer ML to Lisp, you will probably prefer HOL to ZF. + + \medskip The syntax of HOL follows @{text "\"}-calculus and + functional programming. Function application is curried. To apply + the function @{text f} of type @{text "\\<^sub>1 \ \\<^sub>2 \ \\<^sub>3"} to the + arguments @{text a} and @{text b} in HOL, you simply write @{text "f + a b"} (as in ML or Haskell). There is no ``apply'' operator; the + existing application of the Pure @{text "\"}-calculus is re-used. + Note that in HOL @{text "f (a, b)"} means ``@{text "f"} applied to + the pair @{text "(a, b)"} (which is notation for @{text "Pair a + b"}). The latter typically introduces extra formal efforts that can + be avoided by currying functions by default. Explicit tuples are as + infrequent in HOL formalizations as in good ML or Haskell programs. -text {* + \medskip Isabelle/HOL has a distinct feel, compared to other + object-logics like Isabelle/ZF. It identifies object-level types + with meta-level types, taking advantage of the default + type-inference mechanism of Isabelle/Pure. HOL fully identifies + object-level functions with meta-level functions, with native + abstraction and application. + + These identifications allow Isabelle to support HOL particularly + nicely, but they also mean that HOL requires some sophistication + from the user. In particular, an understanding of Hindley-Milner + type-inference with type-classes, which are both used extensively in + the standard libraries and applications. Beginners can set + @{attribute show_types} or even @{attribute show_sorts} to get more + explicit information about the result of type-inference. *} + + +section {* Inductive and coinductive definitions \label{sec:hol-inductive} *} + +text {* An \emph{inductive definition} specifies the least predicate + or set @{text R} closed under given rules: applying a rule to + elements of @{text R} yields a result within @{text R}. For + example, a structural operational semantics is an inductive + definition of an evaluation relation. + + Dually, a \emph{coinductive definition} specifies the greatest + predicate or set @{text R} that is consistent with given rules: + every element of @{text R} can be seen as arising by applying a rule + to elements of @{text R}. An important example is using + bisimulation relations to formalise equivalence of processes and + infinite data structures. + + Both inductive and coinductive definitions are based on the + Knaster-Tarski fixed-point theorem for complete lattices. The + collection of introduction rules given by the user determines a + functor on subsets of set-theoretic relations. The required + monotonicity of the recursion scheme is proven as a prerequisite to + the fixed-point definition and the resulting consequences. This + works by pushing inclusion through logical connectives and any other + operator that might be wrapped around recursive occurrences of the + defined relation: there must be a monotonicity theorem of the form + @{text "A \ B \ \ A \ \ B"}, for each premise @{text "\ R t"} in an + introduction rule. The default rule declarations of Isabelle/HOL + already take care of most common situations. + \begin{matharray}{rcl} - @{command_def (HOL) "typedef"} & : & @{text "local_theory \ proof(prove)"} \\ + @{command_def (HOL) "inductive"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def (HOL) "inductive_set"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def (HOL) "coinductive"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \ local_theory"} \\ + @{attribute_def (HOL) mono} & : & @{text attribute} \\ \end{matharray} @{rail " - @@{command (HOL) typedef} altname? abstype '=' repset + (@@{command (HOL) inductive} | @@{command (HOL) inductive_set} | + @@{command (HOL) coinductive} | @@{command (HOL) coinductive_set}) + @{syntax target}? \\ + @{syntax \"fixes\"} (@'for' @{syntax \"fixes\"})? (@'where' clauses)? \\ + (@'monos' @{syntax thmrefs})? ; - - altname: '(' (@{syntax name} | @'open' | @'open' @{syntax name}) ')' + clauses: (@{syntax thmdecl}? @{syntax prop} + '|') ; - abstype: @{syntax typespec_sorts} @{syntax mixfix}? - ; - repset: @{syntax term} (@'morphisms' @{syntax name} @{syntax name})? + @@{attribute (HOL) mono} (() | 'add' | 'del') "} \begin{description} - \item @{command (HOL) "typedef"}~@{text "(\\<^sub>1, \, \\<^sub>n) t = A"} - axiomatizes a Gordon/HOL-style type definition in the background - theory of the current context, depending on a non-emptiness result - of the set @{text A} (which needs to be proven interactively). + \item @{command (HOL) "inductive"} and @{command (HOL) + "coinductive"} define (co)inductive predicates from the introduction + rules. - The raw type may not depend on parameters or assumptions of the - context --- this is logically impossible in Isabelle/HOL --- but the - non-emptiness property can be local, potentially resulting in - multiple interpretations in target contexts. Thus the established - bijection between the representing set @{text A} and the new type - @{text t} may semantically depend on local assumptions. + The propositions given as @{text "clauses"} in the @{keyword + "where"} part are either rules of the usual @{text "\/\"} format + (with arbitrary nesting), or equalities using @{text "\"}. The + latter specifies extra-logical abbreviations in the sense of + @{command_ref abbreviation}. Introducing abstract syntax + simultaneously with the actual introduction rules is occasionally + useful for complex specifications. - By default, @{command (HOL) "typedef"} defines both a type @{text t} - and a set (term constant) of the same name, unless an alternative - base name is given in parentheses, or the ``@{text "(open)"}'' - declaration is used to suppress a separate constant definition - altogether. The injection from type to set is called @{text Rep_t}, - its inverse @{text Abs_t} --- this may be changed via an explicit - @{keyword (HOL) "morphisms"} declaration. + The optional @{keyword "for"} part contains a list of parameters of + the (co)inductive predicates that remain fixed throughout the + definition, in contrast to arguments of the relation that may vary + in each occurrence within the given @{text "clauses"}. + + The optional @{keyword "monos"} declaration contains additional + \emph{monotonicity theorems}, which are required for each operator + applied to a recursive set in the introduction rules. - Theorems @{text Rep_t}, @{text Rep_t_inverse}, and @{text - Abs_t_inverse} provide the most basic characterization as a - corresponding injection/surjection pair (in both directions). Rules - @{text Rep_t_inject} and @{text Abs_t_inject} provide a slightly - more convenient view on the injectivity part, suitable for automated - proof tools (e.g.\ in @{attribute simp} or @{attribute iff} - declarations). Rules @{text Rep_t_cases}/@{text Rep_t_induct}, and - @{text Abs_t_cases}/@{text Abs_t_induct} provide alternative views - on surjectivity; these are already declared as set or type rules for - the generic @{method cases} and @{method induct} methods. + \item @{command (HOL) "inductive_set"} and @{command (HOL) + "coinductive_set"} are wrappers for to the previous commands for + native HOL predicates. This allows to define (co)inductive sets, + where multiple arguments are simulated via tuples. - An alternative name for the set definition (and other derived - entities) may be specified in parentheses; the default is to use - @{text t} as indicated before. + \item @{attribute (HOL) mono} declares monotonicity rules in the + context. These rule are involved in the automated monotonicity + proof of the above inductive and coinductive definitions. \end{description} *} -section {* Adhoc tuples *} +subsection {* Derived rules *} + +text {* A (co)inductive definition of @{text R} provides the following + main theorems: + + \begin{description} + + \item @{text R.intros} is the list of introduction rules as proven + theorems, for the recursive predicates (or sets). The rules are + also available individually, using the names given them in the + theory file; + + \item @{text R.cases} is the case analysis (or elimination) rule; + + \item @{text R.induct} or @{text R.coinduct} is the (co)induction + rule. + + \end{description} + + When several predicates @{text "R\<^sub>1, \, R\<^sub>n"} are + defined simultaneously, the list of introduction rules is called + @{text "R\<^sub>1_\_R\<^sub>n.intros"}, the case analysis rules are + called @{text "R\<^sub>1.cases, \, R\<^sub>n.cases"}, and the list + of mutual induction rules is called @{text + "R\<^sub>1_\_R\<^sub>n.inducts"}. +*} + + +subsection {* Monotonicity theorems *} + +text {* The context maintains a default set of theorems that are used + in monotonicity proofs. New rules can be declared via the + @{attribute (HOL) mono} attribute. See the main Isabelle/HOL + sources for some examples. The general format of such monotonicity + theorems is as follows: + + \begin{itemize} + + \item Theorems of the form @{text "A \ B \ \ A \ \ B"}, for proving + monotonicity of inductive definitions whose introduction rules have + premises involving terms such as @{text "\ R t"}. + + \item Monotonicity theorems for logical operators, which are of the + general form @{text "(\ \ \) \ \ (\ \ \) \ \ \ \"}. For example, in + the case of the operator @{text "\"}, the corresponding theorem is + \[ + \infer{@{text "P\<^sub>1 \ P\<^sub>2 \ Q\<^sub>1 \ Q\<^sub>2"}}{@{text "P\<^sub>1 \ Q\<^sub>1"} & @{text "P\<^sub>2 \ Q\<^sub>2"}} + \] + + \item De Morgan style equations for reasoning about the ``polarity'' + of expressions, e.g. + \[ + @{prop "\ \ P \ P"} \qquad\qquad + @{prop "\ (P \ Q) \ \ P \ \ Q"} + \] + + \item Equations for reducing complex operators to more primitive + ones whose monotonicity can easily be proved, e.g. + \[ + @{prop "(P \ Q) \ \ P \ Q"} \qquad\qquad + @{prop "Ball A P \ \x. x \ A \ P x"} + \] + + \end{itemize} +*} + +subsubsection {* Examples *} + +text {* The finite powerset operator can be defined inductively like this: *} + +inductive_set Fin :: "'a set \ 'a set set" for A :: "'a set" +where + empty: "{} \ Fin A" +| insert: "a \ A \ B \ Fin A \ insert a B \ Fin A" + +text {* The accessible part of a relation is defined as follows: *} + +inductive acc :: "('a \ 'a \ bool) \ 'a \ bool" + for r :: "'a \ 'a \ bool" (infix "\" 50) +where acc: "(\y. y \ x \ acc r y) \ acc r x" + +text {* Common logical connectives can be easily characterized as +non-recursive inductive definitions with parameters, but without +arguments. *} + +inductive AND for A B :: bool +where "A \ B \ AND A B" + +inductive OR for A B :: bool +where "A \ OR A B" + | "B \ OR A B" + +inductive EXISTS for B :: "'a \ bool" +where "B a \ EXISTS B" + +text {* Here the @{text "cases"} or @{text "induct"} rules produced by + the @{command inductive} package coincide with the expected + elimination rules for Natural Deduction. Already in the original + article by Gerhard Gentzen \cite{Gentzen:1935} there is a hint that + each connective can be characterized by its introductions, and the + elimination can be constructed systematically. *} + + +section {* Recursive functions \label{sec:recursion} *} text {* \begin{matharray}{rcl} - @{attribute_def (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\ + @{command_def (HOL) "primrec"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def (HOL) "fun"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def (HOL) "function"} & : & @{text "local_theory \ proof(prove)"} \\ + @{command_def (HOL) "termination"} & : & @{text "local_theory \ proof(prove)"} \\ \end{matharray} @{rail " - @@{attribute (HOL) split_format} ('(' 'complete' ')')? + @@{command (HOL) primrec} @{syntax target}? @{syntax \"fixes\"} @'where' equations + ; + (@@{command (HOL) fun} | @@{command (HOL) function}) @{syntax target}? functionopts? + @{syntax \"fixes\"} \\ @'where' equations + ; + + equations: (@{syntax thmdecl}? @{syntax prop} + '|') + ; + functionopts: '(' (('sequential' | 'domintros') + ',') ')' + ; + @@{command (HOL) termination} @{syntax term}? "} \begin{description} - \item @{attribute (HOL) split_format}\ @{text "(complete)"} causes - arguments in function applications to be represented canonically - according to their tuple type structure. + \item @{command (HOL) "primrec"} defines primitive recursive + functions over datatypes (see also @{command_ref (HOL) datatype} and + @{command_ref (HOL) rep_datatype}). The given @{text equations} + specify reduction rules that are produced by instantiating the + generic combinator for primitive recursion that is available for + each datatype. + + Each equation needs to be of the form: + + @{text [display] "f x\<^sub>1 \ x\<^sub>m (C y\<^sub>1 \ y\<^sub>k) z\<^sub>1 \ z\<^sub>n = rhs"} + + such that @{text C} is a datatype constructor, @{text rhs} contains + only the free variables on the left-hand side (or from the context), + and all recursive occurrences of @{text "f"} in @{text "rhs"} are of + the form @{text "f \ y\<^sub>i \"} for some @{text i}. At most one + reduction rule for each constructor can be given. The order does + not matter. For missing constructors, the function is defined to + return a default value, but this equation is made difficult to + access for users. + + The reduction rules are declared as @{attribute simp} by default, + which enables standard proof methods like @{method simp} and + @{method auto} to normalize expressions of @{text "f"} applied to + datatype constructions, by simulating symbolic computation via + rewriting. + + \item @{command (HOL) "function"} defines functions by general + wellfounded recursion. A detailed description with examples can be + found in \cite{isabelle-function}. The function is specified by a + set of (possibly conditional) recursive equations with arbitrary + pattern matching. The command generates proof obligations for the + completeness and the compatibility of patterns. + + The defined function is considered partial, and the resulting + simplification rules (named @{text "f.psimps"}) and induction rule + (named @{text "f.pinduct"}) are guarded by a generated domain + predicate @{text "f_dom"}. The @{command (HOL) "termination"} + command can then be used to establish that the function is total. - Note that this operation tends to invent funny names for new local - parameters introduced. + \item @{command (HOL) "fun"} is a shorthand notation for ``@{command + (HOL) "function"}~@{text "(sequential)"}, followed by automated + proof attempts regarding pattern matching and termination. See + \cite{isabelle-function} for further details. + + \item @{command (HOL) "termination"}~@{text f} commences a + termination proof for the previously defined function @{text f}. If + this is omitted, the command refers to the most recent function + definition. After the proof is closed, the recursive equations and + the induction principle is established. + + \end{description} + + Recursive definitions introduced by the @{command (HOL) "function"} + command accommodate reasoning by induction (cf.\ @{method induct}): + rule @{text "f.induct"} refers to a specific induction rule, with + parameters named according to the user-specified equations. Cases + are numbered starting from 1. For @{command (HOL) "primrec"}, the + induction principle coincides with structural recursion on the + datatype where the recursion is carried out. + + The equations provided by these packages may be referred later as + theorem list @{text "f.simps"}, where @{text f} is the (collective) + name of the functions defined. Individual equations may be named + explicitly as well. + + The @{command (HOL) "function"} command accepts the following + options. + + \begin{description} + + \item @{text sequential} enables a preprocessor which disambiguates + overlapping patterns by making them mutually disjoint. Earlier + equations take precedence over later ones. This allows to give the + specification in a format very similar to functional programming. + Note that the resulting simplification and induction rules + correspond to the transformed specification, not the one given + originally. This usually means that each equation given by the user + may result in several theorems. Also note that this automatic + transformation only works for ML-style datatype patterns. + + \item @{text domintros} enables the automated generation of + introduction rules for the domain predicate. While mostly not + needed, they can be helpful in some proofs about partial functions. \end{description} *} +subsubsection {* Example: evaluation of expressions *} + +text {* Subsequently, we define mutual datatypes for arithmetic and + boolean expressions, and use @{command primrec} for evaluation + functions that follow the same recursive structure. *} + +datatype 'a aexp = + IF "'a bexp" "'a aexp" "'a aexp" + | Sum "'a aexp" "'a aexp" + | Diff "'a aexp" "'a aexp" + | Var 'a + | Num nat +and 'a bexp = + Less "'a aexp" "'a aexp" + | And "'a bexp" "'a bexp" + | Neg "'a bexp" + + +text {* \medskip Evaluation of arithmetic and boolean expressions *} + +primrec evala :: "('a \ nat) \ 'a aexp \ nat" + and evalb :: "('a \ nat) \ 'a bexp \ bool" +where + "evala env (IF b a1 a2) = (if evalb env b then evala env a1 else evala env a2)" +| "evala env (Sum a1 a2) = evala env a1 + evala env a2" +| "evala env (Diff a1 a2) = evala env a1 - evala env a2" +| "evala env (Var v) = env v" +| "evala env (Num n) = n" +| "evalb env (Less a1 a2) = (evala env a1 < evala env a2)" +| "evalb env (And b1 b2) = (evalb env b1 \ evalb env b2)" +| "evalb env (Neg b) = (\ evalb env b)" + +text {* Since the value of an expression depends on the value of its + variables, the functions @{const evala} and @{const evalb} take an + additional parameter, an \emph{environment} that maps variables to + their values. + + \medskip Substitution on expressions can be defined similarly. The + mapping @{text f} of type @{typ "'a \ 'a aexp"} given as a + parameter is lifted canonically on the types @{typ "'a aexp"} and + @{typ "'a bexp"}, respectively. +*} + +primrec substa :: "('a \ 'b aexp) \ 'a aexp \ 'b aexp" + and substb :: "('a \ 'b aexp) \ 'a bexp \ 'b bexp" +where + "substa f (IF b a1 a2) = IF (substb f b) (substa f a1) (substa f a2)" +| "substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)" +| "substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)" +| "substa f (Var v) = f v" +| "substa f (Num n) = Num n" +| "substb f (Less a1 a2) = Less (substa f a1) (substa f a2)" +| "substb f (And b1 b2) = And (substb f b1) (substb f b2)" +| "substb f (Neg b) = Neg (substb f b)" + +text {* In textbooks about semantics one often finds substitution + theorems, which express the relationship between substitution and + evaluation. For @{typ "'a aexp"} and @{typ "'a bexp"}, we can prove + such a theorem by mutual induction, followed by simplification. +*} + +lemma subst_one: + "evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a" + "evalb env (substb (Var (v := a')) b) = evalb (env (v := evala env a')) b" + by (induct a and b) simp_all + +lemma subst_all: + "evala env (substa s a) = evala (\x. evala env (s x)) a" + "evalb env (substb s b) = evalb (\x. evala env (s x)) b" + by (induct a and b) simp_all + + +subsubsection {* Example: a substitution function for terms *} + +text {* Functions on datatypes with nested recursion are also defined + by mutual primitive recursion. *} + +datatype ('a, 'b) "term" = Var 'a | App 'b "('a, 'b) term list" + +text {* A substitution function on type @{typ "('a, 'b) term"} can be + defined as follows, by working simultaneously on @{typ "('a, 'b) + term list"}: *} + +primrec subst_term :: "('a \ ('a, 'b) term) \ ('a, 'b) term \ ('a, 'b) term" and + subst_term_list :: "('a \ ('a, 'b) term) \ ('a, 'b) term list \ ('a, 'b) term list" +where + "subst_term f (Var a) = f a" +| "subst_term f (App b ts) = App b (subst_term_list f ts)" +| "subst_term_list f [] = []" +| "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts" + +text {* The recursion scheme follows the structure of the unfolded + definition of type @{typ "('a, 'b) term"}. To prove properties of this + substitution function, mutual induction is needed: +*} + +lemma "subst_term (subst_term f1 \ f2) t = subst_term f1 (subst_term f2 t)" and + "subst_term_list (subst_term f1 \ f2) ts = subst_term_list f1 (subst_term_list f2 ts)" + by (induct t and ts) simp_all + + +subsubsection {* Example: a map function for infinitely branching trees *} + +text {* Defining functions on infinitely branching datatypes by + primitive recursion is just as easy. +*} + +datatype 'a tree = Atom 'a | Branch "nat \ 'a tree" + +primrec map_tree :: "('a \ 'b) \ 'a tree \ 'b tree" +where + "map_tree f (Atom a) = Atom (f a)" +| "map_tree f (Branch ts) = Branch (\x. map_tree f (ts x))" + +text {* Note that all occurrences of functions such as @{text ts} + above must be applied to an argument. In particular, @{term + "map_tree f \ ts"} is not allowed here. *} + +text {* Here is a simple composition lemma for @{term map_tree}: *} + +lemma "map_tree g (map_tree f t) = map_tree (g \ f) t" + by (induct t) simp_all + + +subsection {* Proof methods related to recursive definitions *} + +text {* + \begin{matharray}{rcl} + @{method_def (HOL) pat_completeness} & : & @{text method} \\ + @{method_def (HOL) relation} & : & @{text method} \\ + @{method_def (HOL) lexicographic_order} & : & @{text method} \\ + @{method_def (HOL) size_change} & : & @{text method} \\ + \end{matharray} + + @{rail " + @@{method (HOL) relation} @{syntax term} + ; + @@{method (HOL) lexicographic_order} (@{syntax clasimpmod} * ) + ; + @@{method (HOL) size_change} ( orders (@{syntax clasimpmod} * ) ) + ; + orders: ( 'max' | 'min' | 'ms' ) * + "} + + \begin{description} + + \item @{method (HOL) pat_completeness} is a specialized method to + solve goals regarding the completeness of pattern matching, as + required by the @{command (HOL) "function"} package (cf.\ + \cite{isabelle-function}). + + \item @{method (HOL) relation}~@{text R} introduces a termination + proof using the relation @{text R}. The resulting proof state will + contain goals expressing that @{text R} is wellfounded, and that the + arguments of recursive calls decrease with respect to @{text R}. + Usually, this method is used as the initial proof step of manual + termination proofs. + + \item @{method (HOL) "lexicographic_order"} attempts a fully + automated termination proof by searching for a lexicographic + combination of size measures on the arguments of the function. The + method accepts the same arguments as the @{method auto} method, + which it uses internally to prove local descents. The @{syntax + clasimpmod} modifiers are accepted (as for @{method auto}). + + In case of failure, extensive information is printed, which can help + to analyse the situation (cf.\ \cite{isabelle-function}). + + \item @{method (HOL) "size_change"} also works on termination goals, + using a variation of the size-change principle, together with a + graph decomposition technique (see \cite{krauss_phd} for details). + Three kinds of orders are used internally: @{text max}, @{text min}, + and @{text ms} (multiset), which is only available when the theory + @{text Multiset} is loaded. When no order kinds are given, they are + tried in order. The search for a termination proof uses SAT solving + internally. + + For local descent proofs, the @{syntax clasimpmod} modifiers are + accepted (as for @{method auto}). + + \end{description} +*} + + +subsection {* Functions with explicit partiality *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "partial_function"} & : & @{text "local_theory \ local_theory"} \\ + @{attribute_def (HOL) "partial_function_mono"} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + @@{command (HOL) partial_function} @{syntax target}? + '(' @{syntax nameref} ')' @{syntax \"fixes\"} \\ + @'where' @{syntax thmdecl}? @{syntax prop} + "} + + \begin{description} + + \item @{command (HOL) "partial_function"}~@{text "(mode)"} defines + recursive functions based on fixpoints in complete partial + orders. No termination proof is required from the user or + constructed internally. Instead, the possibility of non-termination + is modelled explicitly in the result type, which contains an + explicit bottom element. + + Pattern matching and mutual recursion are currently not supported. + Thus, the specification consists of a single function described by a + single recursive equation. + + There are no fixed syntactic restrictions on the body of the + function, but the induced functional must be provably monotonic + wrt.\ the underlying order. The monotonicitity proof is performed + internally, and the definition is rejected when it fails. The proof + can be influenced by declaring hints using the + @{attribute (HOL) partial_function_mono} attribute. + + The mandatory @{text mode} argument specifies the mode of operation + of the command, which directly corresponds to a complete partial + order on the result type. By default, the following modes are + defined: + + \begin{description} + \item @{text option} defines functions that map into the @{type + option} type. Here, the value @{term None} is used to model a + non-terminating computation. Monotonicity requires that if @{term + None} is returned by a recursive call, then the overall result + must also be @{term None}. This is best achieved through the use of + the monadic operator @{const "Option.bind"}. + + \item @{text tailrec} defines functions with an arbitrary result + type and uses the slightly degenerated partial order where @{term + "undefined"} is the bottom element. Now, monotonicity requires that + if @{term undefined} is returned by a recursive call, then the + overall result must also be @{term undefined}. In practice, this is + only satisfied when each recursive call is a tail call, whose result + is directly returned. Thus, this mode of operation allows the + definition of arbitrary tail-recursive functions. + \end{description} + + Experienced users may define new modes by instantiating the locale + @{const "partial_function_definitions"} appropriately. + + \item @{attribute (HOL) partial_function_mono} declares rules for + use in the internal monononicity proofs of partial function + definitions. + + \end{description} + +*} + + +subsection {* Old-style recursive function definitions (TFL) *} + +text {* + The old TFL commands @{command (HOL) "recdef"} and @{command (HOL) + "recdef_tc"} for defining recursive are mostly obsolete; @{command + (HOL) "function"} or @{command (HOL) "fun"} should be used instead. + + \begin{matharray}{rcl} + @{command_def (HOL) "recdef"} & : & @{text "theory \ theory)"} \\ + @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \ proof(prove)"} \\ + \end{matharray} + + @{rail " + @@{command (HOL) recdef} ('(' @'permissive' ')')? \\ + @{syntax name} @{syntax term} (@{syntax prop} +) hints? + ; + recdeftc @{syntax thmdecl}? tc + ; + hints: '(' @'hints' ( recdefmod * ) ')' + ; + recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf') + (() | 'add' | 'del') ':' @{syntax thmrefs}) | @{syntax clasimpmod} + ; + tc: @{syntax nameref} ('(' @{syntax nat} ')')? + "} + + \begin{description} + + \item @{command (HOL) "recdef"} defines general well-founded + recursive functions (using the TFL package), see also + \cite{isabelle-HOL}. The ``@{text "(permissive)"}'' option tells + TFL to recover from failed proof attempts, returning unfinished + results. The @{text recdef_simp}, @{text recdef_cong}, and @{text + recdef_wf} hints refer to auxiliary rules to be used in the internal + automated proof process of TFL. Additional @{syntax clasimpmod} + declarations may be given to tune the context of the Simplifier + (cf.\ \secref{sec:simplifier}) and Classical reasoner (cf.\ + \secref{sec:classical}). + + \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the + proof for leftover termination condition number @{text i} (default + 1) as generated by a @{command (HOL) "recdef"} definition of + constant @{text c}. + + Note that in most cases, @{command (HOL) "recdef"} is able to finish + its internal proofs without manual intervention. + + \end{description} + + \medskip Hints for @{command (HOL) "recdef"} may be also declared + globally, using the following attributes. + + \begin{matharray}{rcl} + @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\ + @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\ + @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + (@@{attribute (HOL) recdef_simp} | @@{attribute (HOL) recdef_cong} | + @@{attribute (HOL) recdef_wf}) (() | 'add' | 'del') + "} +*} + + +section {* Datatypes \label{sec:hol-datatype} *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "datatype"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "rep_datatype"} & : & @{text "theory \ proof(prove)"} \\ + \end{matharray} + + @{rail " + @@{command (HOL) datatype} (spec + @'and') + ; + @@{command (HOL) rep_datatype} ('(' (@{syntax name} +) ')')? (@{syntax term} +) + ; + + spec: @{syntax parname}? @{syntax typespec} @{syntax mixfix}? '=' (cons + '|') + ; + cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}? + "} + + \begin{description} + + \item @{command (HOL) "datatype"} defines inductive datatypes in + HOL. + + \item @{command (HOL) "rep_datatype"} represents existing types as + datatypes. + + For foundational reasons, some basic types such as @{typ nat}, @{typ + "'a \ 'b"}, @{typ "'a + 'b"}, @{typ bool} and @{typ unit} are + introduced by more primitive means using @{command_ref typedef}. To + recover the rich infrastructure of @{command datatype} (e.g.\ rules + for @{method cases} and @{method induct} and the primitive recursion + combinators), such types may be represented as actual datatypes + later. This is done by specifying the constructors of the desired + type, and giving a proof of the induction rule, distinctness and + injectivity of constructors. + + For example, see @{file "~~/src/HOL/Sum_Type.thy"} for the + representation of the primitive sum type as fully-featured datatype. + + \end{description} + + The generated rules for @{method induct} and @{method cases} provide + case names according to the given constructors, while parameters are + named after the types (see also \secref{sec:cases-induct}). + + See \cite{isabelle-HOL} for more details on datatypes, but beware of + the old-style theory syntax being used there! Apart from proper + proof methods for case-analysis and induction, there are also + emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL) + induct_tac} available, see \secref{sec:hol-induct-tac}; these admit + to refer directly to the internal structure of subgoals (including + internally bound parameters). +*} + + +subsubsection {* Examples *} + +text {* We define a type of finite sequences, with slightly different + names than the existing @{typ "'a list"} that is already in @{theory + Main}: *} + +datatype 'a seq = Empty | Seq 'a "'a seq" + +text {* We can now prove some simple lemma by structural induction: *} + +lemma "Seq x xs \ xs" +proof (induct xs arbitrary: x) + case Empty + txt {* This case can be proved using the simplifier: the freeness + properties of the datatype are already declared as @{attribute + simp} rules. *} + show "Seq x Empty \ Empty" + by simp +next + case (Seq y ys) + txt {* The step case is proved similarly. *} + show "Seq x (Seq y ys) \ Seq y ys" + using `Seq y ys \ ys` by simp +qed + +text {* Here is a more succinct version of the same proof: *} + +lemma "Seq x xs \ xs" + by (induct xs arbitrary: x) simp_all + section {* Records \label{sec:hol-record} *} @@ -338,48 +1014,184 @@ *} -section {* Datatypes \label{sec:hol-datatype} *} +subsubsection {* Examples *} + +text {* See @{file "~~/src/HOL/ex/Records.thy"}, for example. *} + + +section {* Adhoc tuples *} text {* \begin{matharray}{rcl} - @{command_def (HOL) "datatype"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "rep_datatype"} & : & @{text "theory \ proof(prove)"} \\ + @{attribute_def (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\ \end{matharray} @{rail " - @@{command (HOL) datatype} (spec + @'and') - ; - @@{command (HOL) rep_datatype} ('(' (@{syntax name} +) ')')? (@{syntax term} +) + @@{attribute (HOL) split_format} ('(' 'complete' ')')? + "} + + \begin{description} + + \item @{attribute (HOL) split_format}\ @{text "(complete)"} causes + arguments in function applications to be represented canonically + according to their tuple type structure. + + Note that this operation tends to invent funny names for new local + parameters introduced. + + \end{description} +*} + + +section {* Typedef axiomatization \label{sec:hol-typedef} *} + +text {* A Gordon/HOL-style type definition is a certain axiom scheme + that identifies a new type with a subset of an existing type. More + precisely, the new type is defined by exhibiting an existing type + @{text \}, a set @{text "A :: \ set"}, and a theorem that proves + @{prop "\x. x \ A"}. Thus @{text A} is a non-empty subset of @{text + \}, and the new type denotes this subset. New functions are + postulated that establish an isomorphism between the new type and + the subset. In general, the type @{text \} may involve type + variables @{text "\\<^sub>1, \, \\<^sub>n"} which means that the type definition + produces a type constructor @{text "(\\<^sub>1, \, \\<^sub>n) t"} depending on + those type arguments. + + The axiomatization can be considered a ``definition'' in the sense + of the particular set-theoretic interpretation of HOL + \cite{pitts93}, where the universe of types is required to be + downwards-closed wrt.\ arbitrary non-empty subsets. Thus genuinely + new types introduced by @{command "typedef"} stay within the range + of HOL models by construction. Note that @{command_ref + type_synonym} from Isabelle/Pure merely introduces syntactic + abbreviations, without any logical significance. + + \begin{matharray}{rcl} + @{command_def (HOL) "typedef"} & : & @{text "local_theory \ proof(prove)"} \\ + \end{matharray} + + @{rail " + @@{command (HOL) typedef} alt_name? abs_type '=' rep_set ; - spec: @{syntax parname}? @{syntax typespec} @{syntax mixfix}? '=' (cons + '|') + alt_name: '(' (@{syntax name} | @'open' | @'open' @{syntax name}) ')' ; - cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}? + abs_type: @{syntax typespec_sorts} @{syntax mixfix}? + ; + rep_set: @{syntax term} (@'morphisms' @{syntax name} @{syntax name})? "} \begin{description} - \item @{command (HOL) "datatype"} defines inductive datatypes in - HOL. + \item @{command (HOL) "typedef"}~@{text "(\\<^sub>1, \, \\<^sub>n) t = A"} + axiomatizes a type definition in the background theory of the + current context, depending on a non-emptiness result of the set + @{text A} that needs to be proven here. The set @{text A} may + contain type variables @{text "\\<^sub>1, \, \\<^sub>n"} as specified on the LHS, + but no term variables. + + Even though a local theory specification, the newly introduced type + constructor cannot depend on parameters or assumptions of the + context: this is structurally impossible in HOL. In contrast, the + non-emptiness proof may use local assumptions in unusual situations, + which could result in different interpretations in target contexts: + the meaning of the bijection between the representing set @{text A} + and the new type @{text t} may then change in different application + contexts. + + By default, @{command (HOL) "typedef"} defines both a type + constructor @{text t} for the new type, and a term constant @{text + t} for the representing set within the old type. Use the ``@{text + "(open)"}'' option to suppress a separate constant definition + altogether. The injection from type to set is called @{text Rep_t}, + its inverse @{text Abs_t}, unless explicit @{keyword (HOL) + "morphisms"} specification provides alternative names. - \item @{command (HOL) "rep_datatype"} represents existing types as - inductive ones, generating the standard infrastructure of derived - concepts (primitive recursion etc.). + The core axiomatization uses the locale predicate @{const + type_definition} as defined in Isabelle/HOL. Various basic + consequences of that are instantiated accordingly, re-using the + locale facts with names derived from the new type constructor. Thus + the generic @{thm type_definition.Rep} is turned into the specific + @{text "Rep_t"}, for example. + + Theorems @{thm type_definition.Rep}, @{thm + type_definition.Rep_inverse}, and @{thm type_definition.Abs_inverse} + provide the most basic characterization as a corresponding + injection/surjection pair (in both directions). The derived rules + @{thm type_definition.Rep_inject} and @{thm + type_definition.Abs_inject} provide a more convenient version of + injectivity, suitable for automated proof tools (e.g.\ in + declarations involving @{attribute simp} or @{attribute iff}). + Furthermore, the rules @{thm type_definition.Rep_cases}~/ @{thm + type_definition.Rep_induct}, and @{thm type_definition.Abs_cases}~/ + @{thm type_definition.Abs_induct} provide alternative views on + surjectivity. These rules are already declared as set or type rules + for the generic @{method cases} and @{method induct} methods, + respectively. + + An alternative name for the set definition (and other derived + entities) may be specified in parentheses; the default is to use + @{text t} directly. \end{description} - The induction and exhaustion theorems generated provide case names - according to the constructors involved, while parameters are named - after the types (see also \secref{sec:cases-induct}). + \begin{warn} + If you introduce a new type axiomatically, i.e.\ via @{command_ref + typedecl} and @{command_ref axiomatization}, the minimum requirement + is that it has a non-empty model, to avoid immediate collapse of the + HOL logic. Moreover, one needs to demonstrate that the + interpretation of such free-form axiomatizations can coexist with + that of the regular @{command_def typedef} scheme, and any extension + that other people might have introduced elsewhere (e.g.\ in HOLCF + \cite{MuellerNvOS99}). + \end{warn} +*} + +subsubsection {* Examples *} + +text {* Type definitions permit the introduction of abstract data + types in a safe way, namely by providing models based on already + existing types. Given some abstract axiomatic description @{text P} + of a type, this involves two steps: + + \begin{enumerate} + + \item Find an appropriate type @{text \} and subset @{text A} which + has the desired properties @{text P}, and make a type definition + based on this representation. + + \item Prove that @{text P} holds for @{text \} by lifting @{text P} + from the representation. - See \cite{isabelle-HOL} for more details on datatypes, but beware of - the old-style theory syntax being used there! Apart from proper - proof methods for case-analysis and induction, there are also - emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL) - induct_tac} available, see \secref{sec:hol-induct-tac}; these admit - to refer directly to the internal structure of subgoals (including - internally bound parameters). -*} + \end{enumerate} + + You can later forget about the representation and work solely in + terms of the abstract properties @{text P}. + + \medskip The following trivial example pulls a three-element type + into existence within the formal logical environment of HOL. *} + +typedef three = "{(True, True), (True, False), (False, True)}" + by blast + +definition "One = Abs_three (True, True)" +definition "Two = Abs_three (True, False)" +definition "Three = Abs_three (False, True)" + +lemma three_distinct: "One \ Two" "One \ Three" "Two \ Three" + by (simp_all add: One_def Two_def Three_def Abs_three_inject three_def) + +lemma three_cases: + fixes x :: three obtains "x = One" | "x = Two" | "x = Three" + by (cases x) (auto simp: One_def Two_def Three_def Abs_three_inject three_def) + +text {* Note that such trivial constructions are better done with + derived specification mechanisms such as @{command datatype}: *} + +datatype three' = One' | Two' | Three' + +text {* This avoids re-doing basic definitions and proofs from the + primitive @{command typedef} above. *} section {* Functorial structure of types *} @@ -425,433 +1237,6 @@ *} -section {* Recursive functions \label{sec:recursion} *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "primrec"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def (HOL) "fun"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def (HOL) "function"} & : & @{text "local_theory \ proof(prove)"} \\ - @{command_def (HOL) "termination"} & : & @{text "local_theory \ proof(prove)"} \\ - \end{matharray} - - @{rail " - @@{command (HOL) primrec} @{syntax target}? @{syntax \"fixes\"} @'where' equations - ; - (@@{command (HOL) fun} | @@{command (HOL) function}) @{syntax target}? functionopts? - @{syntax \"fixes\"} \\ @'where' equations - ; - - equations: (@{syntax thmdecl}? @{syntax prop} + '|') - ; - functionopts: '(' (('sequential' | 'domintros') + ',') ')' - ; - @@{command (HOL) termination} @{syntax term}? - "} - - \begin{description} - - \item @{command (HOL) "primrec"} defines primitive recursive - functions over datatypes, see also \cite{isabelle-HOL}. - - \item @{command (HOL) "function"} defines functions by general - wellfounded recursion. A detailed description with examples can be - found in \cite{isabelle-function}. The function is specified by a - set of (possibly conditional) recursive equations with arbitrary - pattern matching. The command generates proof obligations for the - completeness and the compatibility of patterns. - - The defined function is considered partial, and the resulting - simplification rules (named @{text "f.psimps"}) and induction rule - (named @{text "f.pinduct"}) are guarded by a generated domain - predicate @{text "f_dom"}. The @{command (HOL) "termination"} - command can then be used to establish that the function is total. - - \item @{command (HOL) "fun"} is a shorthand notation for ``@{command - (HOL) "function"}~@{text "(sequential)"}, followed by automated - proof attempts regarding pattern matching and termination. See - \cite{isabelle-function} for further details. - - \item @{command (HOL) "termination"}~@{text f} commences a - termination proof for the previously defined function @{text f}. If - this is omitted, the command refers to the most recent function - definition. After the proof is closed, the recursive equations and - the induction principle is established. - - \end{description} - - Recursive definitions introduced by the @{command (HOL) "function"} - command accommodate - reasoning by induction (cf.\ \secref{sec:cases-induct}): rule @{text - "c.induct"} (where @{text c} is the name of the function definition) - refers to a specific induction rule, with parameters named according - to the user-specified equations. Cases are numbered (starting from 1). - - For @{command (HOL) "primrec"}, the induction principle coincides - with structural recursion on the datatype the recursion is carried - out. - - The equations provided by these packages may be referred later as - theorem list @{text "f.simps"}, where @{text f} is the (collective) - name of the functions defined. Individual equations may be named - explicitly as well. - - The @{command (HOL) "function"} command accepts the following - options. - - \begin{description} - - \item @{text sequential} enables a preprocessor which disambiguates - overlapping patterns by making them mutually disjoint. Earlier - equations take precedence over later ones. This allows to give the - specification in a format very similar to functional programming. - Note that the resulting simplification and induction rules - correspond to the transformed specification, not the one given - originally. This usually means that each equation given by the user - may result in several theorems. Also note that this automatic - transformation only works for ML-style datatype patterns. - - \item @{text domintros} enables the automated generation of - introduction rules for the domain predicate. While mostly not - needed, they can be helpful in some proofs about partial functions. - - \end{description} -*} - - -subsection {* Proof methods related to recursive definitions *} - -text {* - \begin{matharray}{rcl} - @{method_def (HOL) pat_completeness} & : & @{text method} \\ - @{method_def (HOL) relation} & : & @{text method} \\ - @{method_def (HOL) lexicographic_order} & : & @{text method} \\ - @{method_def (HOL) size_change} & : & @{text method} \\ - \end{matharray} - - @{rail " - @@{method (HOL) relation} @{syntax term} - ; - @@{method (HOL) lexicographic_order} (@{syntax clasimpmod} * ) - ; - @@{method (HOL) size_change} ( orders (@{syntax clasimpmod} * ) ) - ; - orders: ( 'max' | 'min' | 'ms' ) * - "} - - \begin{description} - - \item @{method (HOL) pat_completeness} is a specialized method to - solve goals regarding the completeness of pattern matching, as - required by the @{command (HOL) "function"} package (cf.\ - \cite{isabelle-function}). - - \item @{method (HOL) relation}~@{text R} introduces a termination - proof using the relation @{text R}. The resulting proof state will - contain goals expressing that @{text R} is wellfounded, and that the - arguments of recursive calls decrease with respect to @{text R}. - Usually, this method is used as the initial proof step of manual - termination proofs. - - \item @{method (HOL) "lexicographic_order"} attempts a fully - automated termination proof by searching for a lexicographic - combination of size measures on the arguments of the function. The - method accepts the same arguments as the @{method auto} method, - which it uses internally to prove local descents. The same context - modifiers as for @{method auto} are accepted, see - \secref{sec:clasimp}. - - In case of failure, extensive information is printed, which can help - to analyse the situation (cf.\ \cite{isabelle-function}). - - \item @{method (HOL) "size_change"} also works on termination goals, - using a variation of the size-change principle, together with a - graph decomposition technique (see \cite{krauss_phd} for details). - Three kinds of orders are used internally: @{text max}, @{text min}, - and @{text ms} (multiset), which is only available when the theory - @{text Multiset} is loaded. When no order kinds are given, they are - tried in order. The search for a termination proof uses SAT solving - internally. - - For local descent proofs, the same context modifiers as for @{method - auto} are accepted, see \secref{sec:clasimp}. - - \end{description} -*} - -subsection {* Functions with explicit partiality *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "partial_function"} & : & @{text "local_theory \ local_theory"} \\ - @{attribute_def (HOL) "partial_function_mono"} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - @@{command (HOL) partial_function} @{syntax target}? - '(' @{syntax nameref} ')' @{syntax \"fixes\"} \\ - @'where' @{syntax thmdecl}? @{syntax prop} - "} - - \begin{description} - - \item @{command (HOL) "partial_function"}~@{text "(mode)"} defines - recursive functions based on fixpoints in complete partial - orders. No termination proof is required from the user or - constructed internally. Instead, the possibility of non-termination - is modelled explicitly in the result type, which contains an - explicit bottom element. - - Pattern matching and mutual recursion are currently not supported. - Thus, the specification consists of a single function described by a - single recursive equation. - - There are no fixed syntactic restrictions on the body of the - function, but the induced functional must be provably monotonic - wrt.\ the underlying order. The monotonicitity proof is performed - internally, and the definition is rejected when it fails. The proof - can be influenced by declaring hints using the - @{attribute (HOL) partial_function_mono} attribute. - - The mandatory @{text mode} argument specifies the mode of operation - of the command, which directly corresponds to a complete partial - order on the result type. By default, the following modes are - defined: - - \begin{description} - \item @{text option} defines functions that map into the @{type - option} type. Here, the value @{term None} is used to model a - non-terminating computation. Monotonicity requires that if @{term - None} is returned by a recursive call, then the overall result - must also be @{term None}. This is best achieved through the use of - the monadic operator @{const "Option.bind"}. - - \item @{text tailrec} defines functions with an arbitrary result - type and uses the slightly degenerated partial order where @{term - "undefined"} is the bottom element. Now, monotonicity requires that - if @{term undefined} is returned by a recursive call, then the - overall result must also be @{term undefined}. In practice, this is - only satisfied when each recursive call is a tail call, whose result - is directly returned. Thus, this mode of operation allows the - definition of arbitrary tail-recursive functions. - \end{description} - - Experienced users may define new modes by instantiating the locale - @{const "partial_function_definitions"} appropriately. - - \item @{attribute (HOL) partial_function_mono} declares rules for - use in the internal monononicity proofs of partial function - definitions. - - \end{description} - -*} - -subsection {* Old-style recursive function definitions (TFL) *} - -text {* - The old TFL commands @{command (HOL) "recdef"} and @{command (HOL) - "recdef_tc"} for defining recursive are mostly obsolete; @{command - (HOL) "function"} or @{command (HOL) "fun"} should be used instead. - - \begin{matharray}{rcl} - @{command_def (HOL) "recdef"} & : & @{text "theory \ theory)"} \\ - @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \ proof(prove)"} \\ - \end{matharray} - - @{rail " - @@{command (HOL) recdef} ('(' @'permissive' ')')? \\ - @{syntax name} @{syntax term} (@{syntax prop} +) hints? - ; - recdeftc @{syntax thmdecl}? tc - ; - hints: '(' @'hints' ( recdefmod * ) ')' - ; - recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf') - (() | 'add' | 'del') ':' @{syntax thmrefs}) | @{syntax clasimpmod} - ; - tc: @{syntax nameref} ('(' @{syntax nat} ')')? - "} - - \begin{description} - - \item @{command (HOL) "recdef"} defines general well-founded - recursive functions (using the TFL package), see also - \cite{isabelle-HOL}. The ``@{text "(permissive)"}'' option tells - TFL to recover from failed proof attempts, returning unfinished - results. The @{text recdef_simp}, @{text recdef_cong}, and @{text - recdef_wf} hints refer to auxiliary rules to be used in the internal - automated proof process of TFL. Additional @{syntax clasimpmod} - declarations (cf.\ \secref{sec:clasimp}) may be given to tune the - context of the Simplifier (cf.\ \secref{sec:simplifier}) and - Classical reasoner (cf.\ \secref{sec:classical}). - - \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the - proof for leftover termination condition number @{text i} (default - 1) as generated by a @{command (HOL) "recdef"} definition of - constant @{text c}. - - Note that in most cases, @{command (HOL) "recdef"} is able to finish - its internal proofs without manual intervention. - - \end{description} - - \medskip Hints for @{command (HOL) "recdef"} may be also declared - globally, using the following attributes. - - \begin{matharray}{rcl} - @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\ - @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\ - @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - (@@{attribute (HOL) recdef_simp} | @@{attribute (HOL) recdef_cong} | - @@{attribute (HOL) recdef_wf}) (() | 'add' | 'del') - "} -*} - - -section {* Inductive and coinductive definitions \label{sec:hol-inductive} *} - -text {* - An \textbf{inductive definition} specifies the least predicate (or - set) @{text R} closed under given rules: applying a rule to elements - of @{text R} yields a result within @{text R}. For example, a - structural operational semantics is an inductive definition of an - evaluation relation. - - Dually, a \textbf{coinductive definition} specifies the greatest - predicate~/ set @{text R} that is consistent with given rules: every - element of @{text R} can be seen as arising by applying a rule to - elements of @{text R}. An important example is using bisimulation - relations to formalise equivalence of processes and infinite data - structures. - - \medskip The HOL package is related to the ZF one, which is - described in a separate paper,\footnote{It appeared in CADE - \cite{paulson-CADE}; a longer version is distributed with Isabelle.} - which you should refer to in case of difficulties. The package is - simpler than that of ZF thanks to implicit type-checking in HOL. - The types of the (co)inductive predicates (or sets) determine the - domain of the fixedpoint definition, and the package does not have - to use inference rules for type-checking. - - \begin{matharray}{rcl} - @{command_def (HOL) "inductive"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def (HOL) "inductive_set"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def (HOL) "coinductive"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \ local_theory"} \\ - @{attribute_def (HOL) mono} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - (@@{command (HOL) inductive} | @@{command (HOL) inductive_set} | - @@{command (HOL) coinductive} | @@{command (HOL) coinductive_set}) - @{syntax target}? @{syntax \"fixes\"} (@'for' @{syntax \"fixes\"})? \\ - (@'where' clauses)? (@'monos' @{syntax thmrefs})? - ; - clauses: (@{syntax thmdecl}? @{syntax prop} + '|') - ; - @@{attribute (HOL) mono} (() | 'add' | 'del') - "} - - \begin{description} - - \item @{command (HOL) "inductive"} and @{command (HOL) - "coinductive"} define (co)inductive predicates from the - introduction rules given in the @{keyword "where"} part. The - optional @{keyword "for"} part contains a list of parameters of the - (co)inductive predicates that remain fixed throughout the - definition. The optional @{keyword "monos"} section contains - \emph{monotonicity theorems}, which are required for each operator - applied to a recursive set in the introduction rules. There - \emph{must} be a theorem of the form @{text "A \ B \ M A \ M B"}, - for each premise @{text "M R\<^sub>i t"} in an introduction rule! - - \item @{command (HOL) "inductive_set"} and @{command (HOL) - "coinductive_set"} are wrappers for to the previous commands, - allowing the definition of (co)inductive sets. - - \item @{attribute (HOL) mono} declares monotonicity rules. These - rule are involved in the automated monotonicity proof of @{command - (HOL) "inductive"}. - - \end{description} -*} - - -subsection {* Derived rules *} - -text {* - Each (co)inductive definition @{text R} adds definitions to the - theory and also proves some theorems: - - \begin{description} - - \item @{text R.intros} is the list of introduction rules as proven - theorems, for the recursive predicates (or sets). The rules are - also available individually, using the names given them in the - theory file; - - \item @{text R.cases} is the case analysis (or elimination) rule; - - \item @{text R.induct} or @{text R.coinduct} is the (co)induction - rule. - - \end{description} - - When several predicates @{text "R\<^sub>1, \, R\<^sub>n"} are - defined simultaneously, the list of introduction rules is called - @{text "R\<^sub>1_\_R\<^sub>n.intros"}, the case analysis rules are - called @{text "R\<^sub>1.cases, \, R\<^sub>n.cases"}, and the list - of mutual induction rules is called @{text - "R\<^sub>1_\_R\<^sub>n.inducts"}. -*} - - -subsection {* Monotonicity theorems *} - -text {* - Each theory contains a default set of theorems that are used in - monotonicity proofs. New rules can be added to this set via the - @{attribute (HOL) mono} attribute. The HOL theory @{text Inductive} - shows how this is done. In general, the following monotonicity - theorems may be added: - - \begin{itemize} - - \item Theorems of the form @{text "A \ B \ M A \ M B"}, for proving - monotonicity of inductive definitions whose introduction rules have - premises involving terms such as @{text "M R\<^sub>i t"}. - - \item Monotonicity theorems for logical operators, which are of the - general form @{text "(\ \ \) \ \ (\ \ \) \ \ \ \"}. For example, in - the case of the operator @{text "\"}, the corresponding theorem is - \[ - \infer{@{text "P\<^sub>1 \ P\<^sub>2 \ Q\<^sub>1 \ Q\<^sub>2"}}{@{text "P\<^sub>1 \ Q\<^sub>1"} & @{text "P\<^sub>2 \ Q\<^sub>2"}} - \] - - \item De Morgan style equations for reasoning about the ``polarity'' - of expressions, e.g. - \[ - @{prop "\ \ P \ P"} \qquad\qquad - @{prop "\ (P \ Q) \ \ P \ \ Q"} - \] - - \item Equations for reducing complex operators to more primitive - ones whose monotonicity can easily be proved, e.g. - \[ - @{prop "(P \ Q) \ \ P \ Q"} \qquad\qquad - @{prop "Ball A P \ \x. x \ A \ P x"} - \] - - \end{itemize} - - %FIXME: Example of an inductive definition -*} - - section {* Arithmetic proof support *} text {* @@ -1021,14 +1406,15 @@ \item @{command (HOL) "value"}~@{text t} evaluates and prints a term; optionally @{text modes} can be specified, which are - appended to the current print mode (see also \cite{isabelle-ref}). + appended to the current print mode; see \secref{sec:print-modes}. Internally, the evaluation is performed by registered evaluators, which are invoked sequentially until a result is returned. Alternatively a specific evaluator can be selected using square brackets; typical evaluators use the current set of code equations - to normalize and include @{text simp} for fully symbolic evaluation - using the simplifier, @{text nbe} for \emph{normalization by evaluation} - and \emph{code} for code generation in SML. + to normalize and include @{text simp} for fully symbolic + evaluation using the simplifier, @{text nbe} for + \emph{normalization by evaluation} and \emph{code} for code + generation in SML. \item @{command (HOL) "quickcheck"} tests the current goal for counterexamples using a series of assignments for its diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/Inner_Syntax.thy --- a/doc-src/IsarRef/Thy/Inner_Syntax.thy Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/Thy/Inner_Syntax.thy Wed Jun 08 10:24:07 2011 +0200 @@ -77,8 +77,8 @@ \end{description} All of the diagnostic commands above admit a list of @{text modes} - to be specified, which is appended to the current print mode (see - also \cite{isabelle-ref}). Thus the output behavior may be modified + to be specified, which is appended to the current print mode; see + \secref{sec:print-modes}. Thus the output behavior may be modified according particular print mode features. For example, @{command "pr"}~@{text "(latex xsymbols)"} would print the current proof state with mathematical symbols and special characters represented in diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/Introduction.thy --- a/doc-src/IsarRef/Thy/Introduction.thy Wed Jun 08 08:47:43 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,73 +0,0 @@ -theory Introduction -imports Base Main -begin - -chapter {* Introduction *} - -section {* Overview *} - -text {* - The \emph{Isabelle} system essentially provides a generic - infrastructure for building deductive systems (programmed in - Standard ML), with a special focus on interactive theorem proving in - higher-order logics. Many years ago, even end-users would refer to - certain ML functions (goal commands, tactics, tacticals etc.) to - pursue their everyday theorem proving tasks. - - In contrast \emph{Isar} provides an interpreted language environment - of its own, which has been specifically tailored for the needs of - theory and proof development. Compared to raw ML, the Isabelle/Isar - top-level provides a more robust and comfortable development - platform, with proper support for theory development graphs, managed - transactions with unlimited undo etc. The Isabelle/Isar version of - the \emph{Proof~General} user interface - \cite{proofgeneral,Aspinall:TACAS:2000} provides a decent front-end - for interactive theory and proof development in this advanced - theorem proving environment, even though it is somewhat biased - towards old-style proof scripts. - - \medskip Apart from the technical advances over bare-bones ML - programming, the main purpose of the Isar language is to provide a - conceptually different view on machine-checked proofs - \cite{Wenzel:1999:TPHOL,Wenzel-PhD}. \emph{Isar} stands for - \emph{Intelligible semi-automated reasoning}. Drawing from both the - traditions of informal mathematical proof texts and high-level - programming languages, Isar offers a versatile environment for - structured formal proof documents. Thus properly written Isar - proofs become accessible to a broader audience than unstructured - tactic scripts (which typically only provide operational information - for the machine). Writing human-readable proof texts certainly - requires some additional efforts by the writer to achieve a good - presentation, both of formal and informal parts of the text. On the - other hand, human-readable formal texts gain some value in their own - right, independently of the mechanic proof-checking process. - - Despite its grand design of structured proof texts, Isar is able to - assimilate the old tactical style as an ``improper'' sub-language. - This provides an easy upgrade path for existing tactic scripts, as - well as some means for interactive experimentation and debugging of - structured proofs. Isabelle/Isar supports a broad range of proof - styles, both readable and unreadable ones. - - \medskip The generic Isabelle/Isar framework (see - \chref{ch:isar-framework}) works reasonably well for any Isabelle - object-logic that conforms to the natural deduction view of the - Isabelle/Pure framework. Specific language elements introduced by - the major object-logics are described in \chref{ch:hol} - (Isabelle/HOL), \chref{ch:holcf} (Isabelle/HOLCF), and \chref{ch:zf} - (Isabelle/ZF). The main language elements are already provided by - the Isabelle/Pure framework. Nevertheless, examples given in the - generic parts will usually refer to Isabelle/HOL as well. - - \medskip Isar commands may be either \emph{proper} document - constructors, or \emph{improper commands}. Some proof methods and - attributes introduced later are classified as improper as well. - Improper Isar language elements, which are marked by ``@{text - "\<^sup>*"}'' in the subsequent chapters; they are often helpful - when developing proof documents, but their use is discouraged for - the final human-readable outcome. Typical examples are diagnostic - commands that print terms or theorems according to the current - context; other commands emulate old-style tactical theorem proving. -*} - -end diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/Preface.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Thy/Preface.thy Wed Jun 08 10:24:07 2011 +0200 @@ -0,0 +1,71 @@ +theory Preface +imports Base Main +begin + +chapter {* Preface *} + +text {* + The \emph{Isabelle} system essentially provides a generic + infrastructure for building deductive systems (programmed in + Standard ML), with a special focus on interactive theorem proving in + higher-order logics. Many years ago, even end-users would refer to + certain ML functions (goal commands, tactics, tacticals etc.) to + pursue their everyday theorem proving tasks. + + In contrast \emph{Isar} provides an interpreted language environment + of its own, which has been specifically tailored for the needs of + theory and proof development. Compared to raw ML, the Isabelle/Isar + top-level provides a more robust and comfortable development + platform, with proper support for theory development graphs, managed + transactions with unlimited undo etc. The Isabelle/Isar version of + the \emph{Proof~General} user interface + \cite{proofgeneral,Aspinall:TACAS:2000} provides a decent front-end + for interactive theory and proof development in this advanced + theorem proving environment, even though it is somewhat biased + towards old-style proof scripts. + + \medskip Apart from the technical advances over bare-bones ML + programming, the main purpose of the Isar language is to provide a + conceptually different view on machine-checked proofs + \cite{Wenzel:1999:TPHOL,Wenzel-PhD}. \emph{Isar} stands for + \emph{Intelligible semi-automated reasoning}. Drawing from both the + traditions of informal mathematical proof texts and high-level + programming languages, Isar offers a versatile environment for + structured formal proof documents. Thus properly written Isar + proofs become accessible to a broader audience than unstructured + tactic scripts (which typically only provide operational information + for the machine). Writing human-readable proof texts certainly + requires some additional efforts by the writer to achieve a good + presentation, both of formal and informal parts of the text. On the + other hand, human-readable formal texts gain some value in their own + right, independently of the mechanic proof-checking process. + + Despite its grand design of structured proof texts, Isar is able to + assimilate the old tactical style as an ``improper'' sub-language. + This provides an easy upgrade path for existing tactic scripts, as + well as some means for interactive experimentation and debugging of + structured proofs. Isabelle/Isar supports a broad range of proof + styles, both readable and unreadable ones. + + \medskip The generic Isabelle/Isar framework (see + \chref{ch:isar-framework}) works reasonably well for any Isabelle + object-logic that conforms to the natural deduction view of the + Isabelle/Pure framework. Specific language elements introduced by + the major object-logics are described in \chref{ch:hol} + (Isabelle/HOL), \chref{ch:holcf} (Isabelle/HOLCF), and \chref{ch:zf} + (Isabelle/ZF). The main language elements are already provided by + the Isabelle/Pure framework. Nevertheless, examples given in the + generic parts will usually refer to Isabelle/HOL as well. + + \medskip Isar commands may be either \emph{proper} document + constructors, or \emph{improper commands}. Some proof methods and + attributes introduced later are classified as improper as well. + Improper Isar language elements, which are marked by ``@{text + "\<^sup>*"}'' in the subsequent chapters; they are often helpful + when developing proof documents, but their use is discouraged for + the final human-readable outcome. Typical examples are diagnostic + commands that print terms or theorems according to the current + context; other commands emulate old-style tactical theorem proving. +*} + +end diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/ROOT.ML --- a/doc-src/IsarRef/Thy/ROOT.ML Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/Thy/ROOT.ML Wed Jun 08 10:24:07 2011 +0200 @@ -1,7 +1,8 @@ quick_and_dirty := true; use_thys [ - "Introduction", + "Preface", + "Synopsis", "Framework", "First_Order_Logic", "Outer_Syntax", diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/Spec.thy --- a/doc-src/IsarRef/Thy/Spec.thy Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/Thy/Spec.thy Wed Jun 08 10:24:07 2011 +0200 @@ -2,7 +2,7 @@ imports Base Main begin -chapter {* Theory specifications *} +chapter {* Specifications *} text {* The Isabelle/Isar theory format integrates specifications and @@ -922,14 +922,14 @@ Thm.rule_attribute (fn context: Context.generic => fn th: thm => let val th' = th OF ths - in th' end)) *} "my rule" + in th' end)) *} attribute_setup my_declaration = {* Attrib.thms >> (fn ths => Thm.declaration_attribute (fn th: thm => fn context: Context.generic => let val context' = context - in context' end)) *} "my declaration" + in context' end)) *} section {* Primitive specification elements *} diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/Synopsis.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Thy/Synopsis.thy Wed Jun 08 10:24:07 2011 +0200 @@ -0,0 +1,1108 @@ +theory Synopsis +imports Base Main +begin + +chapter {* Synopsis *} + +section {* Notepad *} + +text {* + An Isar proof body serves as mathematical notepad to compose logical + content, consisting of types, terms, facts. +*} + + +subsection {* Types and terms *} + +notepad +begin + txt {* Locally fixed entities: *} + fix x -- {* local constant, without any type information yet *} + fix x :: 'a -- {* variant with explicit type-constraint for subsequent use*} + + fix a b + assume "a = b" -- {* type assignment at first occurrence in concrete term *} + + txt {* Definitions (non-polymorphic): *} + def x \ "t::'a" + + txt {* Abbreviations (polymorphic): *} + let ?f = "\x. x" + term "?f ?f" + + txt {* Notation: *} + write x ("***") +end + + +subsection {* Facts *} + +text {* + A fact is a simultaneous list of theorems. +*} + + +subsubsection {* Producing facts *} + +notepad +begin + + txt {* Via assumption (``lambda''): *} + assume a: A + + txt {* Via proof (``let''): *} + have b: B sorry + + txt {* Via abbreviation (``let''): *} + note c = a b + +end + + +subsubsection {* Referencing facts *} + +notepad +begin + txt {* Via explicit name: *} + assume a: A + note a + + txt {* Via implicit name: *} + assume A + note this + + txt {* Via literal proposition (unification with results from the proof text): *} + assume A + note `A` + + assume "\x. B x" + note `B a` + note `B b` +end + + +subsubsection {* Manipulating facts *} + +notepad +begin + txt {* Instantiation: *} + assume a: "\x. B x" + note a + note a [of b] + note a [where x = b] + + txt {* Backchaining: *} + assume 1: A + assume 2: "A \ C" + note 2 [OF 1] + note 1 [THEN 2] + + txt {* Symmetric results: *} + assume "x = y" + note this [symmetric] + + assume "x \ y" + note this [symmetric] + + txt {* Adhoc-simplification (take care!): *} + assume "P ([] @ xs)" + note this [simplified] +end + + +subsubsection {* Projections *} + +text {* + Isar facts consist of multiple theorems. There is notation to project + interval ranges. +*} + +notepad +begin + assume stuff: A B C D + note stuff(1) + note stuff(2-3) + note stuff(2-) +end + + +subsubsection {* Naming conventions *} + +text {* + \begin{itemize} + + \item Lower-case identifiers are usually preferred. + + \item Facts can be named after the main term within the proposition. + + \item Facts should \emph{not} be named after the command that + introduced them (@{command "assume"}, @{command "have"}). This is + misleading and hard to maintain. + + \item Natural numbers can be used as ``meaningless'' names (more + appropriate than @{text "a1"}, @{text "a2"} etc.) + + \item Symbolic identifiers are supported (e.g. @{text "*"}, @{text + "**"}, @{text "***"}). + + \end{itemize} +*} + + +subsection {* Block structure *} + +text {* + The formal notepad is block structured. The fact produced by the last + entry of a block is exported into the outer context. +*} + +notepad +begin + { + have a: A sorry + have b: B sorry + note a b + } + note this + note `A` + note `B` +end + +text {* Explicit blocks as well as implicit blocks of nested goal + statements (e.g.\ @{command have}) automatically introduce one extra + pair of parentheses in reserve. The @{command next} command allows + to ``jump'' between these sub-blocks. *} + +notepad +begin + + { + have a: A sorry + next + have b: B + proof - + show B sorry + next + have c: C sorry + next + have d: D sorry + qed + } + + txt {* Alternative version with explicit parentheses everywhere: *} + + { + { + have a: A sorry + } + { + have b: B + proof - + { + show B sorry + } + { + have c: C sorry + } + { + have d: D sorry + } + qed + } + } + +end + + +section {* Calculational reasoning \label{sec:calculations-synopsis} *} + +text {* + For example, see @{file "~~/src/HOL/Isar_Examples/Group.thy"}. +*} + + +subsection {* Special names in Isar proofs *} + +text {* + \begin{itemize} + + \item term @{text "?thesis"} --- the main conclusion of the + innermost pending claim + + \item term @{text "\"} --- the argument of the last explicitly + stated result (for infix application this is the right-hand side) + + \item fact @{text "this"} --- the last result produced in the text + + \end{itemize} +*} + +notepad +begin + have "x = y" + proof - + term ?thesis + show ?thesis sorry + term ?thesis -- {* static! *} + qed + term "\" + thm this +end + +text {* Calculational reasoning maintains the special fact called + ``@{text calculation}'' in the background. Certain language + elements combine primary @{text this} with secondary @{text + calculation}. *} + + +subsection {* Transitive chains *} + +text {* The Idea is to combine @{text this} and @{text calculation} + via typical @{text trans} rules (see also @{command + print_trans_rules}): *} + +thm trans +thm less_trans +thm less_le_trans + +notepad +begin + txt {* Plain bottom-up calculation: *} + have "a = b" sorry + also + have "b = c" sorry + also + have "c = d" sorry + finally + have "a = d" . + + txt {* Variant using the @{text "\"} abbreviation: *} + have "a = b" sorry + also + have "\ = c" sorry + also + have "\ = d" sorry + finally + have "a = d" . + + txt {* Top-down version with explicit claim at the head: *} + have "a = d" + proof - + have "a = b" sorry + also + have "\ = c" sorry + also + have "\ = d" sorry + finally + show ?thesis . + qed +next + txt {* Mixed inequalities (require suitable base type): *} + fix a b c d :: nat + + have "a < b" sorry + also + have "b\ c" sorry + also + have "c = d" sorry + finally + have "a < d" . +end + + +subsubsection {* Notes *} + +text {* + \begin{itemize} + + \item The notion of @{text trans} rule is very general due to the + flexibility of Isabelle/Pure rule composition. + + \item User applications may declare there own rules, with some care + about the operational details of higher-order unification. + + \end{itemize} +*} + + +subsection {* Degenerate calculations and bigstep reasoning *} + +text {* The Idea is to append @{text this} to @{text calculation}, + without rule composition. *} + +notepad +begin + txt {* A vacuous proof: *} + have A sorry + moreover + have B sorry + moreover + have C sorry + ultimately + have A and B and C . +next + txt {* Slightly more content (trivial bigstep reasoning): *} + have A sorry + moreover + have B sorry + moreover + have C sorry + ultimately + have "A \ B \ C" by blast +next + txt {* More ambitious bigstep reasoning involving structured results: *} + have "A \ B \ C" sorry + moreover + { assume A have R sorry } + moreover + { assume B have R sorry } + moreover + { assume C have R sorry } + ultimately + have R by blast -- {* ``big-bang integration'' of proof blocks (occasionally fragile) *} +end + + +section {* Induction *} + +subsection {* Induction as Natural Deduction *} + +text {* In principle, induction is just a special case of Natural + Deduction (see also \secref{sec:natural-deduction-synopsis}). For + example: *} + +thm nat.induct +print_statement nat.induct + +notepad +begin + fix n :: nat + have "P n" + proof (rule nat.induct) -- {* fragile rule application! *} + show "P 0" sorry + next + fix n :: nat + assume "P n" + show "P (Suc n)" sorry + qed +end + +text {* + In practice, much more proof infrastructure is required. + + The proof method @{method induct} provides: + \begin{itemize} + + \item implicit rule selection and robust instantiation + + \item context elements via symbolic case names + + \item support for rule-structured induction statements, with local + parameters, premises, etc. + + \end{itemize} +*} + +notepad +begin + fix n :: nat + have "P n" + proof (induct n) + case 0 + show ?case sorry + next + case (Suc n) + from Suc.hyps show ?case sorry + qed +end + + +subsubsection {* Example *} + +text {* + The subsequent example combines the following proof patterns: + \begin{itemize} + + \item outermost induction (over the datatype structure of natural + numbers), to decompose the proof problem in top-down manner + + \item calculational reasoning (\secref{sec:calculations-synopsis}) + to compose the result in each case + + \item solving local claims within the calculation by simplification + + \end{itemize} +*} + +lemma + fixes n :: nat + shows "(\i=0..n. i) = n * (n + 1) div 2" +proof (induct n) + case 0 + have "(\i=0..0. i) = (0::nat)" by simp + also have "\ = 0 * (0 + 1) div 2" by simp + finally show ?case . +next + case (Suc n) + have "(\i=0..Suc n. i) = (\i=0..n. i) + (n + 1)" by simp + also have "\ = n * (n + 1) div 2 + (n + 1)" by (simp add: Suc.hyps) + also have "\ = (n * (n + 1) + 2 * (n + 1)) div 2" by simp + also have "\ = (Suc n * (Suc n + 1)) div 2" by simp + finally show ?case . +qed + +text {* This demonstrates how induction proofs can be done without + having to consider the raw Natural Deduction structure. *} + + +subsection {* Induction with local parameters and premises *} + +text {* Idea: Pure rule statements are passed through the induction + rule. This achieves convenient proof patterns, thanks to some + internal trickery in the @{method induct} method. + + Important: Using compact HOL formulae with @{text "\/\"} is a + well-known anti-pattern! It would produce useless formal noise. +*} + +notepad +begin + fix n :: nat + fix P :: "nat \ bool" + fix Q :: "'a \ nat \ bool" + + have "P n" + proof (induct n) + case 0 + show "P 0" sorry + next + case (Suc n) + from `P n` show "P (Suc n)" sorry + qed + + have "A n \ P n" + proof (induct n) + case 0 + from `A 0` show "P 0" sorry + next + case (Suc n) + from `A n \ P n` + and `A (Suc n)` show "P (Suc n)" sorry + qed + + have "\x. Q x n" + proof (induct n) + case 0 + show "Q x 0" sorry + next + case (Suc n) + from `\x. Q x n` show "Q x (Suc n)" sorry + txt {* Local quantification admits arbitrary instances: *} + note `Q a n` and `Q b n` + qed +end + + +subsection {* Implicit induction context *} + +text {* The @{method induct} method can isolate local parameters and + premises directly from the given statement. This is convenient in + practical applications, but requires some understanding of what is + going on internally (as explained above). *} + +notepad +begin + fix n :: nat + fix Q :: "'a \ nat \ bool" + + fix x :: 'a + assume "A x n" + then have "Q x n" + proof (induct n arbitrary: x) + case 0 + from `A x 0` show "Q x 0" sorry + next + case (Suc n) + from `\x. A x n \ Q x n` -- {* arbitrary instances can be produced here *} + and `A x (Suc n)` show "Q x (Suc n)" sorry + qed +end + + +subsection {* Advanced induction with term definitions *} + +text {* Induction over subexpressions of a certain shape are delicate + to formalize. The Isar @{method induct} method provides + infrastructure for this. + + Idea: sub-expressions of the problem are turned into a defined + induction variable; often accompanied with fixing of auxiliary + parameters in the original expression. *} + +notepad +begin + fix a :: "'a \ nat" + fix A :: "nat \ bool" + + assume "A (a x)" + then have "P (a x)" + proof (induct "a x" arbitrary: x) + case 0 + note prem = `A (a x)` + and defn = `0 = a x` + show "P (a x)" sorry + next + case (Suc n) + note hyp = `\x. n = a x \ A (a x) \ P (a x)` + and prem = `A (a x)` + and defn = `Suc n = a x` + show "P (a x)" sorry + qed +end + + +section {* Natural Deduction \label{sec:natural-deduction-synopsis} *} + +subsection {* Rule statements *} + +text {* + Isabelle/Pure ``theorems'' are always natural deduction rules, + which sometimes happen to consist of a conclusion only. + + The framework connectives @{text "\"} and @{text "\"} indicate the + rule structure declaratively. For example: *} + +thm conjI +thm impI +thm nat.induct + +text {* + The object-logic is embedded into the Pure framework via an implicit + derivability judgment @{term "Trueprop :: bool \ prop"}. + + Thus any HOL formulae appears atomic to the Pure framework, while + the rule structure outlines the corresponding proof pattern. + + This can be made explicit as follows: +*} + +notepad +begin + write Trueprop ("Tr") + + thm conjI + thm impI + thm nat.induct +end + +text {* + Isar provides first-class notation for rule statements as follows. +*} + +print_statement conjI +print_statement impI +print_statement nat.induct + + +subsubsection {* Examples *} + +text {* + Introductions and eliminations of some standard connectives of + the object-logic can be written as rule statements as follows. (The + proof ``@{command "by"}~@{method blast}'' serves as sanity check.) +*} + +lemma "(P \ False) \ \ P" by blast +lemma "\ P \ P \ Q" by blast + +lemma "P \ Q \ P \ Q" by blast +lemma "P \ Q \ (P \ Q \ R) \ R" by blast + +lemma "P \ P \ Q" by blast +lemma "Q \ P \ Q" by blast +lemma "P \ Q \ (P \ R) \ (Q \ R) \ R" by blast + +lemma "(\x. P x) \ (\x. P x)" by blast +lemma "(\x. P x) \ P x" by blast + +lemma "P x \ (\x. P x)" by blast +lemma "(\x. P x) \ (\x. P x \ R) \ R" by blast + +lemma "x \ A \ x \ B \ x \ A \ B" by blast +lemma "x \ A \ B \ (x \ A \ x \ B \ R) \ R" by blast + +lemma "x \ A \ x \ A \ B" by blast +lemma "x \ B \ x \ A \ B" by blast +lemma "x \ A \ B \ (x \ A \ R) \ (x \ B \ R) \ R" by blast + + +subsection {* Isar context elements *} + +text {* We derive some results out of the blue, using Isar context + elements and some explicit blocks. This illustrates their meaning + wrt.\ Pure connectives, without goal states getting in the way. *} + +notepad +begin + { + fix x + have "B x" sorry + } + have "\x. B x" by fact + +next + + { + assume A + have B sorry + } + have "A \ B" by fact + +next + + { + def x \ t + have "B x" sorry + } + have "B t" by fact + +next + + { + obtain x :: 'a where "B x" sorry + have C sorry + } + have C by fact + +end + + +subsection {* Pure rule composition *} + +text {* + The Pure framework provides means for: + + \begin{itemize} + + \item backward-chaining of rules by @{inference resolution} + + \item closing of branches by @{inference assumption} + + \end{itemize} + + Both principles involve higher-order unification of @{text \}-terms + modulo @{text "\\\"}-equivalence (cf.\ Huet and Miller). *} + +notepad +begin + assume a: A and b: B + thm conjI + thm conjI [of A B] -- "instantiation" + thm conjI [of A B, OF a b] -- "instantiation and composition" + thm conjI [OF a b] -- "composition via unification (trivial)" + thm conjI [OF `A` `B`] + + thm conjI [OF disjI1] +end + +text {* Note: Low-level rule composition is tedious and leads to + unreadable~/ unmaintainable expressions in the text. *} + + +subsection {* Structured backward reasoning *} + +text {* Idea: Canonical proof decomposition via @{command fix}~/ + @{command assume}~/ @{command show}, where the body produces a + natural deduction rule to refine some goal. *} + +notepad +begin + fix A B :: "'a \ bool" + + have "\x. A x \ B x" + proof - + fix x + assume "A x" + show "B x" sorry + qed + + have "\x. A x \ B x" + proof - + { + fix x + assume "A x" + show "B x" sorry + } -- "implicit block structure made explicit" + note `\x. A x \ B x` + -- "side exit for the resulting rule" + qed +end + + +subsection {* Structured rule application *} + +text {* + Idea: Previous facts and new claims are composed with a rule from + the context (or background library). +*} + +notepad +begin + assume r1: "A \ B \ C" -- {* simple rule (Horn clause) *} + + have A sorry -- "prefix of facts via outer sub-proof" + then have C + proof (rule r1) + show B sorry -- "remaining rule premises via inner sub-proof" + qed + + have C + proof (rule r1) + show A sorry + show B sorry + qed + + have A and B sorry + then have C + proof (rule r1) + qed + + have A and B sorry + then have C by (rule r1) + +next + + assume r2: "A \ (\x. B1 x \ B2 x) \ C" -- {* nested rule *} + + have A sorry + then have C + proof (rule r2) + fix x + assume "B1 x" + show "B2 x" sorry + qed + + txt {* The compound rule premise @{prop "\x. B1 x \ B2 x"} is better + addressed via @{command fix}~/ @{command assume}~/ @{command show} + in the nested proof body. *} +end + + +subsection {* Example: predicate logic *} + +text {* + Using the above principles, standard introduction and elimination proofs + of predicate logic connectives of HOL work as follows. +*} + +notepad +begin + have "A \ B" and A sorry + then have B .. + + have A sorry + then have "A \ B" .. + + have B sorry + then have "A \ B" .. + + have "A \ B" sorry + then have C + proof + assume A + then show C sorry + next + assume B + then show C sorry + qed + + have A and B sorry + then have "A \ B" .. + + have "A \ B" sorry + then have A .. + + have "A \ B" sorry + then have B .. + + have False sorry + then have A .. + + have True .. + + have "\ A" + proof + assume A + then show False sorry + qed + + have "\ A" and A sorry + then have B .. + + have "\x. P x" + proof + fix x + show "P x" sorry + qed + + have "\x. P x" sorry + then have "P a" .. + + have "\x. P x" + proof + show "P a" sorry + qed + + have "\x. P x" sorry + then have C + proof + fix a + assume "P a" + show C sorry + qed + + txt {* Less awkward version using @{command obtain}: *} + have "\x. P x" sorry + then obtain a where "P a" .. +end + +text {* Further variations to illustrate Isar sub-proofs involving + @{command show}: *} + +notepad +begin + have "A \ B" + proof -- {* two strictly isolated subproofs *} + show A sorry + next + show B sorry + qed + + have "A \ B" + proof -- {* one simultaneous sub-proof *} + show A and B sorry + qed + + have "A \ B" + proof -- {* two subproofs in the same context *} + show A sorry + show B sorry + qed + + have "A \ B" + proof -- {* swapped order *} + show B sorry + show A sorry + qed + + have "A \ B" + proof -- {* sequential subproofs *} + show A sorry + show B using `A` sorry + qed +end + + +subsubsection {* Example: set-theoretic operators *} + +text {* There is nothing special about logical connectives (@{text + "\"}, @{text "\"}, @{text "\"}, @{text "\"} etc.). Operators from + set-theory or lattice-theory for analogously. It is only a matter + of rule declarations in the library; rules can be also specified + explicitly. +*} + +notepad +begin + have "x \ A" and "x \ B" sorry + then have "x \ A \ B" .. + + have "x \ A" sorry + then have "x \ A \ B" .. + + have "x \ B" sorry + then have "x \ A \ B" .. + + have "x \ A \ B" sorry + then have C + proof + assume "x \ A" + then show C sorry + next + assume "x \ B" + then show C sorry + qed + +next + have "x \ \A" + proof + fix a + assume "a \ A" + show "x \ a" sorry + qed + + have "x \ \A" sorry + then have "x \ a" + proof + show "a \ A" sorry + qed + + have "a \ A" and "x \ a" sorry + then have "x \ \A" .. + + have "x \ \A" sorry + then obtain a where "a \ A" and "x \ a" .. +end + + +section {* Generalized elimination and cases *} + +subsection {* General elimination rules *} + +text {* + The general format of elimination rules is illustrated by the + following typical representatives: +*} + +thm exE -- {* local parameter *} +thm conjE -- {* local premises *} +thm disjE -- {* split into cases *} + +text {* + Combining these characteristics leads to the following general scheme + for elimination rules with cases: + + \begin{itemize} + + \item prefix of assumptions (or ``major premises'') + + \item one or more cases that enable to establish the main conclusion + in an augmented context + + \end{itemize} +*} + +notepad +begin + assume r: + "A1 \ A2 \ (* assumptions *) + (\x y. B1 x y \ C1 x y \ R) \ (* case 1 *) + (\x y. B2 x y \ C2 x y \ R) \ (* case 2 *) + R (* main conclusion *)" + + have A1 and A2 sorry + then have R + proof (rule r) + fix x y + assume "B1 x y" and "C1 x y" + show ?thesis sorry + next + fix x y + assume "B2 x y" and "C2 x y" + show ?thesis sorry + qed +end + +text {* Here @{text "?thesis"} is used to refer to the unchanged goal + statement. *} + + +subsection {* Rules with cases *} + +text {* + Applying an elimination rule to some goal, leaves that unchanged + but allows to augment the context in the sub-proof of each case. + + Isar provides some infrastructure to support this: + + \begin{itemize} + + \item native language elements to state eliminations + + \item symbolic case names + + \item method @{method cases} to recover this structure in a + sub-proof + + \end{itemize} +*} + +print_statement exE +print_statement conjE +print_statement disjE + +lemma + assumes A1 and A2 -- {* assumptions *} + obtains + (case1) x y where "B1 x y" and "C1 x y" + | (case2) x y where "B2 x y" and "C2 x y" + sorry + + +subsubsection {* Example *} + +lemma tertium_non_datur: + obtains + (T) A + | (F) "\ A" + by blast + +notepad +begin + fix x y :: 'a + have C + proof (cases "x = y" rule: tertium_non_datur) + case T + from `x = y` show ?thesis sorry + next + case F + from `x \ y` show ?thesis sorry + qed +end + + +subsubsection {* Example *} + +text {* + Isabelle/HOL specification mechanisms (datatype, inductive, etc.) + provide suitable derived cases rules. +*} + +datatype foo = Foo | Bar foo + +notepad +begin + fix x :: foo + have C + proof (cases x) + case Foo + from `x = Foo` show ?thesis sorry + next + case (Bar a) + from `x = Bar a` show ?thesis sorry + qed +end + + +subsection {* Obtaining local contexts *} + +text {* A single ``case'' branch may be inlined into Isar proof text + via @{command obtain}. This proves @{prop "(\x. B x \ thesis) \ + thesis"} on the spot, and augments the context afterwards. *} + +notepad +begin + fix B :: "'a \ bool" + + obtain x where "B x" sorry + note `B x` + + txt {* Conclusions from this context may not mention @{term x} again! *} + { + obtain x where "B x" sorry + from `B x` have C sorry + } + note `C` +end + +end \ No newline at end of file diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/document/Document_Preparation.tex --- a/doc-src/IsarRef/Thy/document/Document_Preparation.tex Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/Thy/document/Document_Preparation.tex Wed Jun 08 10:24:07 2011 +0200 @@ -446,7 +446,7 @@ \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsubsubsection{Styled antiquotations% +\isamarkupsubsection{Styled antiquotations% } \isamarkuptrue% % @@ -476,7 +476,7 @@ \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsubsubsection{General options% +\isamarkupsubsection{General options% } \isamarkuptrue% % diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/document/Generic.tex --- a/doc-src/IsarRef/Thy/document/Generic.tex Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/Thy/document/Generic.tex Wed Jun 08 10:24:07 2011 +0200 @@ -734,7 +734,23 @@ \isamarkuptrue% % \begin{isamarkuptext}% -\begin{matharray}{rcl} +Simplification procedures are ML functions that produce proven + rewrite rules on demand. They are associated with higher-order + patterns that approximate the left-hand sides of equations. The + Simplifier first matches the current redex against one of the LHS + patterns; if this succeeds, the corresponding ML function is + invoked, passing the Simplifier context and redex term. Thus rules + may be specifically fashioned for particular situations, resulting + in a more powerful mechanism than term rewriting by a fixed set of + rules. + + Any successful result needs to be a (possibly conditional) rewrite + rule \isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ u{\isaliteral{22}{\isachardoublequote}}} that is applicable to the current redex. The + rule will be applied just as any ordinary rewrite rule. It is + expected to be already in \emph{internal form}, bypassing the + automatic preprocessing of object-level equivalences. + + \begin{matharray}{rcl} \indexdef{}{command}{simproc\_setup}\hypertarget{command.simproc-setup}{\hyperlink{command.simproc-setup}{\mbox{\isa{\isacommand{simproc{\isaliteral{5F}{\isacharunderscore}}setup}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ simproc & : & \isa{attribute} \\ \end{matharray} @@ -810,6 +826,48 @@ \end{isamarkuptext}% \isamarkuptrue% % +\isamarkupsubsubsection{Example% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The following simplification procedure for \isa{{\isaliteral{28}{\isacharparenleft}}{\isaliteral{3F}{\isacharquery}}u{\isaliteral{5C3C436F6C6F6E3E}{\isasymColon}}unit{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{29}{\isacharparenright}}} in HOL performs fine-grained + control over rule application, beyond higher-order pattern matching. + Declaring \isa{unit{\isaliteral{5F}{\isacharunderscore}}eq} as \hyperlink{attribute.simp}{\mbox{\isa{simp}}} directly would make + the simplifier loop! Note that a version of this simplification + procedure is already active in Isabelle/HOL.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{simproc{\isaliteral{5F}{\isacharunderscore}}setup}\isamarkupfalse% +\ unit\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}x{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}unit{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline +\ \ fn\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ ct\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline +\ \ \ \ if\ HOLogic{\isaliteral{2E}{\isachardot}}is{\isaliteral{5F}{\isacharunderscore}}unit\ {\isaliteral{28}{\isacharparenleft}}term{\isaliteral{5F}{\isacharunderscore}}of\ ct{\isaliteral{29}{\isacharparenright}}\ then\ NONE\isanewline +\ \ \ \ else\ SOME\ {\isaliteral{28}{\isacharparenleft}}mk{\isaliteral{5F}{\isacharunderscore}}meta{\isaliteral{5F}{\isacharunderscore}}eq\ % +\isaantiq +thm\ unit{\isaliteral{5F}{\isacharunderscore}}eq{}% +\endisaantiq +{\isaliteral{29}{\isacharparenright}}\isanewline +{\isaliteral{2A7D}{\isacharverbatimclose}}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +Since the Simplifier applies simplification procedures + frequently, it is important to make the failure check in ML + reasonably fast.% +\end{isamarkuptext}% +\isamarkuptrue% +% \isamarkupsubsection{Forward simplification% } \isamarkuptrue% @@ -866,30 +924,334 @@ } \isamarkuptrue% % -\isamarkupsubsection{Basic methods% +\isamarkupsubsection{Basic concepts% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Although Isabelle is generic, many users will be working in + some extension of classical first-order logic. Isabelle/ZF is built + upon theory FOL, while Isabelle/HOL conceptually contains + first-order logic as a fragment. Theorem-proving in predicate logic + is undecidable, but many automated strategies have been developed to + assist in this task. + + Isabelle's classical reasoner is a generic package that accepts + certain information about a logic and delivers a suite of automatic + proof tools, based on rules that are classified and declared in the + context. These proof procedures are slow and simplistic compared + with high-end automated theorem provers, but they can save + considerable time and effort in practice. They can prove theorems + such as Pelletier's \cite{pelletier86} problems 40 and 41 in a few + milliseconds (including full proof reconstruction):% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ F\ x\ y\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ F\ x\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}z{\isaliteral{2E}{\isachardot}}\ F\ z\ y\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ F\ z\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}z{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ y\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ f\ x\ z\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ f\ x\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}z{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +The proof tools are generic. They are not restricted to + first-order logic, and have been heavily used in the development of + the Isabelle/HOL library and applications. The tactics can be + traced, and their components can be called directly; in this manner, + any proof can be viewed interactively.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{The sequent calculus% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle supports natural deduction, which is easy to use for + interactive proof. But natural deduction does not easily lend + itself to automation, and has a bias towards intuitionism. For + certain proofs in classical logic, it can not be called natural. + The \emph{sequent calculus}, a generalization of natural deduction, + is easier to automate. + + A \textbf{sequent} has the form \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{22}{\isachardoublequote}}}, where \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}{\isaliteral{22}{\isachardoublequote}}} + and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{22}{\isachardoublequote}}} are sets of formulae.\footnote{For first-order + logic, sequents can equivalently be made from lists or multisets of + formulae.} The sequent \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} is + \textbf{valid} if \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}} implies \isa{{\isaliteral{22}{\isachardoublequote}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}. Thus \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}} represent assumptions, each of which + is true, while \isa{{\isaliteral{22}{\isachardoublequote}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} represent alternative goals. A + sequent is \textbf{basic} if its left and right sides have a common + formula, as in \isa{{\isaliteral{22}{\isachardoublequote}}P{\isaliteral{2C}{\isacharcomma}}\ Q\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ Q{\isaliteral{2C}{\isacharcomma}}\ R{\isaliteral{22}{\isachardoublequote}}}; basic sequents are trivially + valid. + + Sequent rules are classified as \textbf{right} or \textbf{left}, + indicating which side of the \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}{\isaliteral{22}{\isachardoublequote}}} symbol they operate on. + Rules that operate on the right side are analogous to natural + deduction's introduction rules, and left rules are analogous to + elimination rules. The sequent calculus analogue of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}I{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} + is the rule + \[ + \infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}P{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ Q{\isaliteral{22}{\isachardoublequote}}}} + \] + Applying the rule backwards, this breaks down some implication on + the right side of a sequent; \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{22}{\isachardoublequote}}} stand for + the sets of formulae that are unaffected by the inference. The + analogue of the pair \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}I{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}I{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} is the + single rule + \[ + \infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ P{\isaliteral{2C}{\isacharcomma}}\ Q{\isaliteral{22}{\isachardoublequote}}}} + \] + This breaks down some disjunction on the right side, replacing it by + both disjuncts. Thus, the sequent calculus is a kind of + multiple-conclusion logic. + + To illustrate the use of multiple formulae on the right, let us + prove the classical theorem \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{28}{\isacharparenleft}}Q\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. Working + backwards, we reduce this formula to a basic sequent: + \[ + \infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{28}{\isacharparenleft}}Q\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}} + {\infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}Q\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}} + {\infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}P\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ Q{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}Q\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}} + {\isa{{\isaliteral{22}{\isachardoublequote}}P{\isaliteral{2C}{\isacharcomma}}\ Q\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ Q{\isaliteral{2C}{\isacharcomma}}\ P{\isaliteral{22}{\isachardoublequote}}}}}} + \] + + This example is typical of the sequent calculus: start with the + desired theorem and apply rules backwards in a fairly arbitrary + manner. This yields a surprisingly effective proof procedure. + Quantifiers add only few complications, since Isabelle handles + parameters and schematic variables. See \cite[Chapter + 10]{paulson-ml2} for further discussion.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Simulating sequents by natural deduction% } \isamarkuptrue% % \begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{method}{rule}\hypertarget{method.rule}{\hyperlink{method.rule}{\mbox{\isa{rule}}}} & : & \isa{method} \\ - \indexdef{}{method}{contradiction}\hypertarget{method.contradiction}{\hyperlink{method.contradiction}{\mbox{\isa{contradiction}}}} & : & \isa{method} \\ - \indexdef{}{method}{intro}\hypertarget{method.intro}{\hyperlink{method.intro}{\mbox{\isa{intro}}}} & : & \isa{method} \\ - \indexdef{}{method}{elim}\hypertarget{method.elim}{\hyperlink{method.elim}{\mbox{\isa{elim}}}} & : & \isa{method} \\ +Isabelle can represent sequents directly, as in the + object-logic LK. But natural deduction is easier to work with, and + most object-logics employ it. Fortunately, we can simulate the + sequent \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} by the Isabelle formula + \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} where the order of + the assumptions and the choice of \isa{{\isaliteral{22}{\isachardoublequote}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} are arbitrary. + Elim-resolution plays a key role in simulating sequent proofs. + + We can easily handle reasoning on the left. Elim-resolution with + the rules \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} achieves + a similar effect as the corresponding sequent rules. For the other + connectives, we use sequent-style elimination rules instead of + destruction rules such as \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616E643E}{\isasymand}}E{\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. + But note that the rule \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}L{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} has no effect under our + representation of sequents! + \[ + \infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}L{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ P{\isaliteral{22}{\isachardoublequote}}}} + \] + + What about reasoning on the right? Introduction rules can only + affect the formula in the conclusion, namely \isa{{\isaliteral{22}{\isachardoublequote}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}. The + other right-side formulae are represented as negated assumptions, + \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}. In order to operate on one of these, it + must first be exchanged with \isa{{\isaliteral{22}{\isachardoublequote}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}. Elim-resolution with the + \isa{swap} rule has this effect: \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ R\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequote}}} + + To ensure that swaps occur only when necessary, each introduction + rule is converted into a swapped form: it is resolved with the + second premise of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}swap{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. The swapped form of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616E643E}{\isasymand}}I{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, which might be called \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}{\isaliteral{5C3C616E643E}{\isasymand}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, is + \begin{isabelle}% +{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C616E643E}{\isasymand}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ R\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ R\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequote}}% +\end{isabelle} + + Similarly, the swapped form of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}I{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} is + \begin{isabelle}% +{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ R\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequote}}% +\end{isabelle} + + Swapped introduction rules are applied using elim-resolution, which + deletes the negated formula. Our representation of sequents also + requires the use of ordinary introduction rules. If we had no + regard for readability of intermediate goal states, we could treat + the right side more uniformly by representing sequents as \begin{isabelle}% +{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}{\isaliteral{22}{\isachardoublequote}}% +\end{isabelle}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Extra rules for the sequent calculus% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +As mentioned, destruction rules such as \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616E643E}{\isasymand}}E{\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and + \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} must be replaced by sequent-style elimination rules. + In addition, we need rules to embody the classical equivalence + between \isa{{\isaliteral{22}{\isachardoublequote}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}}. The introduction + rules \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}I{\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} are replaced by a rule that simulates + \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}: \begin{isabelle}% +{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}% +\end{isabelle} + + The destruction rule \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} is replaced by \begin{isabelle}% +{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequote}}% +\end{isabelle} + + Quantifier replication also requires special rules. In classical + logic, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}} is equivalent to \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}; + the rules \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}L{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} are dual: + \[ + \infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{2C}{\isacharcomma}}\ P\ t{\isaliteral{22}{\isachardoublequote}}}} + \qquad + \infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}L{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}P\ t{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{22}{\isachardoublequote}}}} + \] + Thus both kinds of quantifier may be replicated. Theorems requiring + multiple uses of a universal formula are easy to invent; consider + \begin{isabelle}% +{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P\ {\isaliteral{28}{\isacharparenleft}}f\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ P\ a\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P\ {\isaliteral{28}{\isacharparenleft}}f\isaliteral{5C3C5E7375703E}{}\isactrlsup n\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}% +\end{isabelle} for any + \isa{{\isaliteral{22}{\isachardoublequote}}n\ {\isaliteral{3E}{\isachargreater}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}. Natural examples of the multiple use of an + existential formula are rare; a standard one is \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}y{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P\ y{\isaliteral{22}{\isachardoublequote}}}. + + Forgoing quantifier replication loses completeness, but gains + decidability, since the search space becomes finite. Many useful + theorems can be proved without replication, and the search generally + delivers its verdict in a reasonable time. To adopt this approach, + represent the sequent rules \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}L{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and + \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} by \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}I{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}I{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, + respectively, and put \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} into elimination form: \begin{isabelle}% +{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}P\ t\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequote}}% +\end{isabelle} + + Elim-resolution with this rule will delete the universal formula + after a single use. To replicate universal quantifiers, replace the + rule by \begin{isabelle}% +{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}P\ t\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequote}}% +\end{isabelle} + + To replicate existential quantifiers, replace \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}I{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} by + \begin{isabelle}% +{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ t{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}% +\end{isabelle} + + All introduction rules mentioned above are also useful in swapped + form. + + Replication makes the search space infinite; we must apply the rules + with care. The classical reasoner distinguishes between safe and + unsafe rules, applying the latter only when there is no alternative. + Depth-first search may well go down a blind alley; best-first search + is better behaved in an infinite search space. However, quantifier + replication is too expensive to prove any but the simplest theorems.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Rule declarations% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The proof tools of the Classical Reasoner depend on + collections of rules declared in the context, which are classified + as introduction, elimination or destruction and as \emph{safe} or + \emph{unsafe}. In general, safe rules can be attempted blindly, + while unsafe rules must be used with care. A safe rule must never + reduce a provable goal to an unprovable set of subgoals. + + The rule \isa{{\isaliteral{22}{\isachardoublequote}}P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}} is unsafe because it reduces \isa{{\isaliteral{22}{\isachardoublequote}}P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}} to \isa{{\isaliteral{22}{\isachardoublequote}}P{\isaliteral{22}{\isachardoublequote}}}, which might turn out as premature choice of an + unprovable subgoal. Any rule is unsafe whose premises contain new + unknowns. The elimination rule \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}P\ t\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequote}}} is + unsafe, since it is applied via elim-resolution, which discards the + assumption \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}} and replaces it by the weaker + assumption \isa{{\isaliteral{22}{\isachardoublequote}}P\ t{\isaliteral{22}{\isachardoublequote}}}. The rule \isa{{\isaliteral{22}{\isachardoublequote}}P\ t\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}} is + unsafe for similar reasons. The quantifier duplication rule \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}P\ t\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequote}}} is unsafe in a different sense: + since it keeps the assumption \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}}, it is prone to + looping. In classical first-order logic, all rules are safe except + those mentioned above. + + The safe~/ unsafe distinction is vague, and may be regarded merely + as a way of giving some rules priority over others. One could argue + that \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} is unsafe, because repeated application of it + could generate exponentially many subgoals. Induction rules are + unsafe because inductive proofs are difficult to set up + automatically. Any inference is unsafe that instantiates an unknown + in the proof state --- thus matching must be used, rather than + unification. Even proof by assumption is unsafe if it instantiates + unknowns shared with other subgoals. + + \begin{matharray}{rcl} + \indexdef{}{command}{print\_claset}\hypertarget{command.print-claset}{\hyperlink{command.print-claset}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}claset}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{}{attribute}{intro}\hypertarget{attribute.intro}{\hyperlink{attribute.intro}{\mbox{\isa{intro}}}} & : & \isa{attribute} \\ + \indexdef{}{attribute}{elim}\hypertarget{attribute.elim}{\hyperlink{attribute.elim}{\mbox{\isa{elim}}}} & : & \isa{attribute} \\ + \indexdef{}{attribute}{dest}\hypertarget{attribute.dest}{\hyperlink{attribute.dest}{\mbox{\isa{dest}}}} & : & \isa{attribute} \\ + \indexdef{}{attribute}{rule}\hypertarget{attribute.rule}{\hyperlink{attribute.rule}{\mbox{\isa{rule}}}} & : & \isa{attribute} \\ + \indexdef{}{attribute}{iff}\hypertarget{attribute.iff}{\hyperlink{attribute.iff}{\mbox{\isa{iff}}}} & : & \isa{attribute} \\ + \indexdef{}{attribute}{swapped}\hypertarget{attribute.swapped}{\hyperlink{attribute.swapped}{\mbox{\isa{swapped}}}} & : & \isa{attribute} \\ \end{matharray} \begin{railoutput} \rail@begin{3}{} \rail@bar -\rail@term{\hyperlink{method.rule}{\mbox{\isa{rule}}}}[] +\rail@term{\hyperlink{attribute.intro}{\mbox{\isa{intro}}}}[] \rail@nextbar{1} -\rail@term{\hyperlink{method.intro}{\mbox{\isa{intro}}}}[] +\rail@term{\hyperlink{attribute.elim}{\mbox{\isa{elim}}}}[] \rail@nextbar{2} -\rail@term{\hyperlink{method.elim}{\mbox{\isa{elim}}}}[] +\rail@term{\hyperlink{attribute.dest}{\mbox{\isa{dest}}}}[] +\rail@endbar +\rail@bar +\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[] +\rail@nextbar{1} +\rail@nextbar{2} +\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] \rail@endbar \rail@bar \rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] +\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] +\rail@endbar +\rail@end +\rail@begin{1}{} +\rail@term{\hyperlink{attribute.rule}{\mbox{\isa{rule}}}}[] +\rail@term{\isa{del}}[] +\rail@end +\rail@begin{3}{} +\rail@term{\hyperlink{attribute.iff}{\mbox{\isa{iff}}}}[] +\rail@bar +\rail@bar +\rail@nextbar{1} +\rail@term{\isa{add}}[] +\rail@endbar +\rail@bar +\rail@nextbar{1} +\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] +\rail@endbar +\rail@nextbar{2} +\rail@term{\isa{del}}[] \rail@endbar \rail@end \end{railoutput} @@ -897,27 +1259,53 @@ \begin{description} - \item \hyperlink{method.rule}{\mbox{\isa{rule}}} as offered by the Classical Reasoner is a - refinement over the primitive one (see \secref{sec:pure-meth-att}). - Both versions essentially work the same, but the classical version - observes the classical rule context in addition to that of - Isabelle/Pure. + \item \hyperlink{command.print-claset}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}claset}}}} prints the collection of rules + declared to the Classical Reasoner, i.e.\ the \verb|claset| + within the context. - Common object logics (HOL, ZF, etc.) declare a rich collection of - classical rules (even if these would qualify as intuitionistic - ones), but only few declarations to the rule context of - Isabelle/Pure (\secref{sec:pure-meth-att}). + \item \hyperlink{attribute.intro}{\mbox{\isa{intro}}}, \hyperlink{attribute.elim}{\mbox{\isa{elim}}}, and \hyperlink{attribute.dest}{\mbox{\isa{dest}}} + declare introduction, elimination, and destruction rules, + respectively. By default, rules are considered as \emph{unsafe} + (i.e.\ not applied blindly without backtracking), while ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}}'' classifies as \emph{safe}. Rule declarations marked by + ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'' coincide with those of Isabelle/Pure, cf.\ + \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps + of the \hyperlink{method.rule}{\mbox{\isa{rule}}} method). The optional natural number + specifies an explicit weight argument, which is ignored by the + automated reasoning tools, but determines the search order of single + rule steps. + + Introduction rules are those that can be applied using ordinary + resolution. Their swapped forms are generated internally, which + will be applied using elim-resolution. Elimination rules are + applied using elim-resolution. Rules are sorted by the number of + new subgoals they will yield; rules that generate the fewest + subgoals will be tried first. Otherwise, later declarations take + precedence over earlier ones. - \item \hyperlink{method.contradiction}{\mbox{\isa{contradiction}}} solves some goal by contradiction, - deriving any result from both \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequote}}} and \isa{A}. Chained - facts, which are guaranteed to participate, may appear in either - order. + Rules already present in the context with the same classification + are ignored. A warning is printed if the rule has already been + added with some other classification, but the rule is added anyway + as requested. + + \item \hyperlink{attribute.rule}{\mbox{\isa{rule}}}~\isa{del} deletes all occurrences of a + rule from the classical context, regardless of its classification as + introduction~/ elimination~/ destruction and safe~/ unsafe. - \item \hyperlink{method.intro}{\mbox{\isa{intro}}} and \hyperlink{method.elim}{\mbox{\isa{elim}}} repeatedly refine some goal - by intro- or elim-resolution, after having inserted any chained - facts. Exactly the rules given as arguments are taken into account; - this allows fine-tuned decomposition of a proof problem, in contrast - to common automated tools. + \item \hyperlink{attribute.iff}{\mbox{\isa{iff}}} declares logical equivalences to the + Simplifier and the Classical reasoner at the same time. + Non-conditional rules result in a safe introduction and elimination + pair; conditional ones are considered unsafe. Rules with negative + conclusion are automatically inverted (using \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}{\isaliteral{22}{\isachardoublequote}}}-elimination + internally). + + The ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'' version of \hyperlink{attribute.iff}{\mbox{\isa{iff}}} declares rules to + the Isabelle/Pure context only, and omits the Simplifier + declaration. + + \item \hyperlink{attribute.swapped}{\mbox{\isa{swapped}}} turns an introduction rule into an + elimination, by resolving with the classical swap principle \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ R\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequote}}} in the second position. This is mainly for + illustrative purposes: the Classical Reasoner already swaps rules + internally as explained above. \end{description}% \end{isamarkuptext}% @@ -930,11 +1318,14 @@ \begin{isamarkuptext}% \begin{matharray}{rcl} \indexdef{}{method}{blast}\hypertarget{method.blast}{\hyperlink{method.blast}{\mbox{\isa{blast}}}} & : & \isa{method} \\ + \indexdef{}{method}{auto}\hypertarget{method.auto}{\hyperlink{method.auto}{\mbox{\isa{auto}}}} & : & \isa{method} \\ + \indexdef{}{method}{force}\hypertarget{method.force}{\hyperlink{method.force}{\mbox{\isa{force}}}} & : & \isa{method} \\ \indexdef{}{method}{fast}\hypertarget{method.fast}{\hyperlink{method.fast}{\mbox{\isa{fast}}}} & : & \isa{method} \\ \indexdef{}{method}{slow}\hypertarget{method.slow}{\hyperlink{method.slow}{\mbox{\isa{slow}}}} & : & \isa{method} \\ \indexdef{}{method}{best}\hypertarget{method.best}{\hyperlink{method.best}{\mbox{\isa{best}}}} & : & \isa{method} \\ - \indexdef{}{method}{safe}\hypertarget{method.safe}{\hyperlink{method.safe}{\mbox{\isa{safe}}}} & : & \isa{method} \\ - \indexdef{}{method}{clarify}\hypertarget{method.clarify}{\hyperlink{method.clarify}{\mbox{\isa{clarify}}}} & : & \isa{method} \\ + \indexdef{}{method}{fastsimp}\hypertarget{method.fastsimp}{\hyperlink{method.fastsimp}{\mbox{\isa{fastsimp}}}} & : & \isa{method} \\ + \indexdef{}{method}{slowsimp}\hypertarget{method.slowsimp}{\hyperlink{method.slowsimp}{\mbox{\isa{slowsimp}}}} & : & \isa{method} \\ + \indexdef{}{method}{bestsimp}\hypertarget{method.bestsimp}{\hyperlink{method.bestsimp}{\mbox{\isa{bestsimp}}}} & : & \isa{method} \\ \end{matharray} \begin{railoutput} @@ -949,23 +1340,51 @@ \rail@cnont{\hyperlink{syntax.clamod}{\mbox{\isa{clamod}}}}[] \rail@endplus \rail@end -\rail@begin{5}{} +\rail@begin{2}{} +\rail@term{\hyperlink{method.auto}{\mbox{\isa{auto}}}}[] +\rail@bar +\rail@nextbar{1} +\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] +\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] +\rail@endbar +\rail@plus +\rail@nextplus{1} +\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] +\rail@endplus +\rail@end +\rail@begin{2}{} +\rail@term{\hyperlink{method.force}{\mbox{\isa{force}}}}[] +\rail@plus +\rail@nextplus{1} +\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] +\rail@endplus +\rail@end +\rail@begin{3}{} \rail@bar \rail@term{\hyperlink{method.fast}{\mbox{\isa{fast}}}}[] \rail@nextbar{1} \rail@term{\hyperlink{method.slow}{\mbox{\isa{slow}}}}[] \rail@nextbar{2} \rail@term{\hyperlink{method.best}{\mbox{\isa{best}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{method.safe}{\mbox{\isa{safe}}}}[] -\rail@nextbar{4} -\rail@term{\hyperlink{method.clarify}{\mbox{\isa{clarify}}}}[] \rail@endbar \rail@plus \rail@nextplus{1} \rail@cnont{\hyperlink{syntax.clamod}{\mbox{\isa{clamod}}}}[] \rail@endplus \rail@end +\rail@begin{3}{} +\rail@bar +\rail@term{\hyperlink{method.fastsimp}{\mbox{\isa{fastsimp}}}}[] +\rail@nextbar{1} +\rail@term{\hyperlink{method.slowsimp}{\mbox{\isa{slowsimp}}}}[] +\rail@nextbar{2} +\rail@term{\hyperlink{method.bestsimp}{\mbox{\isa{bestsimp}}}}[] +\rail@endbar +\rail@plus +\rail@nextplus{1} +\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] +\rail@endplus +\rail@end \rail@begin{4}{\indexdef{}{syntax}{clamod}\hypertarget{syntax.clamod}{\hyperlink{syntax.clamod}{\mbox{\isa{clamod}}}}} \rail@bar \rail@bar @@ -987,72 +1406,6 @@ \rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] \rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] \rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.blast}{\mbox{\isa{blast}}} refers to the classical tableau prover (see - \verb|blast_tac| in \cite{isabelle-ref}). The optional argument - specifies a user-supplied search bound (default 20). - - \item \hyperlink{method.fast}{\mbox{\isa{fast}}}, \hyperlink{method.slow}{\mbox{\isa{slow}}}, \hyperlink{method.best}{\mbox{\isa{best}}}, \hyperlink{method.safe}{\mbox{\isa{safe}}}, and \hyperlink{method.clarify}{\mbox{\isa{clarify}}} refer to the generic classical - reasoner. See \verb|fast_tac|, \verb|slow_tac|, \verb|best_tac|, \verb|safe_tac|, and \verb|clarify_tac| in \cite{isabelle-ref} for more - information. - - \end{description} - - Any of the above methods support additional modifiers of the context - of classical rules. Their semantics is analogous to the attributes - given before. Facts provided by forward chaining are inserted into - the goal before commencing proof search.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Combined automated methods \label{sec:clasimp}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{method}{auto}\hypertarget{method.auto}{\hyperlink{method.auto}{\mbox{\isa{auto}}}} & : & \isa{method} \\ - \indexdef{}{method}{force}\hypertarget{method.force}{\hyperlink{method.force}{\mbox{\isa{force}}}} & : & \isa{method} \\ - \indexdef{}{method}{clarsimp}\hypertarget{method.clarsimp}{\hyperlink{method.clarsimp}{\mbox{\isa{clarsimp}}}} & : & \isa{method} \\ - \indexdef{}{method}{fastsimp}\hypertarget{method.fastsimp}{\hyperlink{method.fastsimp}{\mbox{\isa{fastsimp}}}} & : & \isa{method} \\ - \indexdef{}{method}{slowsimp}\hypertarget{method.slowsimp}{\hyperlink{method.slowsimp}{\mbox{\isa{slowsimp}}}} & : & \isa{method} \\ - \indexdef{}{method}{bestsimp}\hypertarget{method.bestsimp}{\hyperlink{method.bestsimp}{\mbox{\isa{bestsimp}}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{method.auto}{\mbox{\isa{auto}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] -\rail@endplus -\rail@end -\rail@begin{5}{} -\rail@bar -\rail@term{\hyperlink{method.force}{\mbox{\isa{force}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{method.clarsimp}{\mbox{\isa{clarsimp}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{method.fastsimp}{\mbox{\isa{fastsimp}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{method.slowsimp}{\mbox{\isa{slowsimp}}}}[] -\rail@nextbar{4} -\rail@term{\hyperlink{method.bestsimp}{\mbox{\isa{bestsimp}}}}[] -\rail@endbar -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] -\rail@endplus -\rail@end \rail@begin{14}{\indexdef{}{syntax}{clasimpmod}\hypertarget{syntax.clasimpmod}{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}} \rail@bar \rail@term{\isa{simp}}[] @@ -1117,73 +1470,166 @@ \begin{description} - \item \hyperlink{method.auto}{\mbox{\isa{auto}}}, \hyperlink{method.force}{\mbox{\isa{force}}}, \hyperlink{method.clarsimp}{\mbox{\isa{clarsimp}}}, \hyperlink{method.fastsimp}{\mbox{\isa{fastsimp}}}, \hyperlink{method.slowsimp}{\mbox{\isa{slowsimp}}}, and \hyperlink{method.bestsimp}{\mbox{\isa{bestsimp}}} provide access - to Isabelle's combined simplification and classical reasoning - tactics. These correspond to \verb|auto_tac|, \verb|force_tac|, \verb|clarsimp_tac|, and Classical Reasoner tactics with the Simplifier - added as wrapper, see \cite{isabelle-ref} for more information. The - modifier arguments correspond to those given in - \secref{sec:simplifier} and \secref{sec:classical}. Just note that - the ones related to the Simplifier are prefixed by \isa{simp} - here. + \item \hyperlink{method.blast}{\mbox{\isa{blast}}} is a separate classical tableau prover that + uses the same classical rule declarations as explained before. + + Proof search is coded directly in ML using special data structures. + A successful proof is then reconstructed using regular Isabelle + inferences. It is faster and more powerful than the other classical + reasoning tools, but has major limitations too. + + \begin{itemize} + + \item It does not use the classical wrapper tacticals, such as the + integration with the Simplifier of \hyperlink{method.fastsimp}{\mbox{\isa{fastsimp}}}. + + \item It does not perform higher-order unification, as needed by the + rule \isa{{\isaliteral{3F}{\isacharquery}}f\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ range\ {\isaliteral{3F}{\isacharquery}}f} in HOL. There are often + alternatives to such rules, for example \isa{{\isaliteral{3F}{\isacharquery}}b\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}f\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}b\ {\isaliteral{5C3C696E3E}{\isasymin}}\ range\ {\isaliteral{3F}{\isacharquery}}f}. + + \item Function variables may only be applied to parameters of the + subgoal. (This restriction arises because the prover does not use + higher-order unification.) If other function variables are present + then the prover will fail with the message \texttt{Function Var's + argument not a bound variable}. + + \item Its proof strategy is more general than \hyperlink{method.fast}{\mbox{\isa{fast}}} but can + be slower. If \hyperlink{method.blast}{\mbox{\isa{blast}}} fails or seems to be running forever, + try \hyperlink{method.fast}{\mbox{\isa{fast}}} and the other proof tools described below. + + \end{itemize} + + The optional integer argument specifies a bound for the number of + unsafe steps used in a proof. By default, \hyperlink{method.blast}{\mbox{\isa{blast}}} starts + with a bound of 0 and increases it successively to 20. In contrast, + \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}blast\ lim{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} tries to prove the goal using a search bound + of \isa{{\isaliteral{22}{\isachardoublequote}}lim{\isaliteral{22}{\isachardoublequote}}}. Sometimes a slow proof using \hyperlink{method.blast}{\mbox{\isa{blast}}} can + be made much faster by supplying the successful search bound to this + proof method instead. + + \item \hyperlink{method.auto}{\mbox{\isa{auto}}} combines classical reasoning with + simplification. It is intended for situations where there are a lot + of mostly trivial subgoals; it proves all the easy ones, leaving the + ones it cannot prove. Occasionally, attempting to prove the hard + ones may take a long time. + + %FIXME auto nat arguments + + \item \hyperlink{method.force}{\mbox{\isa{force}}} is intended to prove the first subgoal + completely, using many fancy proof tools and performing a rather + exhaustive search. As a result, proof attempts may take rather long + or diverge easily. + + \item \hyperlink{method.fast}{\mbox{\isa{fast}}}, \hyperlink{method.best}{\mbox{\isa{best}}}, \hyperlink{method.slow}{\mbox{\isa{slow}}} attempt to + prove the first subgoal using sequent-style reasoning as explained + before. Unlike \hyperlink{method.blast}{\mbox{\isa{blast}}}, they construct proofs directly in + Isabelle. + + There is a difference in search strategy and back-tracking: \hyperlink{method.fast}{\mbox{\isa{fast}}} uses depth-first search and \hyperlink{method.best}{\mbox{\isa{best}}} uses best-first + search (guided by a heuristic function: normally the total size of + the proof state). + + Method \hyperlink{method.slow}{\mbox{\isa{slow}}} is like \hyperlink{method.fast}{\mbox{\isa{fast}}}, but conducts a broader + search: it may, when backtracking from a failed proof attempt, undo + even the step of proving a subgoal by assumption. - Facts provided by forward chaining are inserted into the goal before - doing the search. + \item \hyperlink{method.fastsimp}{\mbox{\isa{fastsimp}}}, \hyperlink{method.slowsimp}{\mbox{\isa{slowsimp}}}, \hyperlink{method.bestsimp}{\mbox{\isa{bestsimp}}} are + like \hyperlink{method.fast}{\mbox{\isa{fast}}}, \hyperlink{method.slow}{\mbox{\isa{slow}}}, \hyperlink{method.best}{\mbox{\isa{best}}}, respectively, + but use the Simplifier as additional wrapper. + + \end{description} + + Any of the above methods support additional modifiers of the context + of classical (and simplifier) rules, but the ones related to the + Simplifier are explicitly prefixed by \isa{simp} here. The + semantics of these ad-hoc rule declarations is analogous to the + attributes given before. Facts provided by forward chaining are + inserted into the goal before commencing proof search.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Semi-automated methods% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +These proof methods may help in situations when the + fully-automated tools fail. The result is a simpler subgoal that + can be tackled by other means, such as by manual instantiation of + quantifiers. + + \begin{matharray}{rcl} + \indexdef{}{method}{safe}\hypertarget{method.safe}{\hyperlink{method.safe}{\mbox{\isa{safe}}}} & : & \isa{method} \\ + \indexdef{}{method}{clarify}\hypertarget{method.clarify}{\hyperlink{method.clarify}{\mbox{\isa{clarify}}}} & : & \isa{method} \\ + \indexdef{}{method}{clarsimp}\hypertarget{method.clarsimp}{\hyperlink{method.clarsimp}{\mbox{\isa{clarsimp}}}} & : & \isa{method} \\ + \end{matharray} + + \begin{railoutput} +\rail@begin{2}{} +\rail@bar +\rail@term{\hyperlink{method.safe}{\mbox{\isa{safe}}}}[] +\rail@nextbar{1} +\rail@term{\hyperlink{method.clarify}{\mbox{\isa{clarify}}}}[] +\rail@endbar +\rail@plus +\rail@nextplus{1} +\rail@cnont{\hyperlink{syntax.clamod}{\mbox{\isa{clamod}}}}[] +\rail@endplus +\rail@end +\rail@begin{2}{} +\rail@term{\hyperlink{method.clarsimp}{\mbox{\isa{clarsimp}}}}[] +\rail@plus +\rail@nextplus{1} +\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] +\rail@endplus +\rail@end +\end{railoutput} + + + \begin{description} + + \item \hyperlink{method.safe}{\mbox{\isa{safe}}} repeatedly performs safe steps on all subgoals. + It is deterministic, with at most one outcome. + + \item \hyperlink{method.clarify}{\mbox{\isa{clarify}}} performs a series of safe steps as follows. + + No splitting step is applied; for example, the subgoal \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequote}}} is left as a conjunction. Proof by assumption, Modus Ponens, + etc., may be performed provided they do not instantiate unknowns. + Assumptions of the form \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3D}{\isacharequal}}\ t{\isaliteral{22}{\isachardoublequote}}} may be eliminated. The safe + wrapper tactical is applied. + + \item \hyperlink{method.clarsimp}{\mbox{\isa{clarsimp}}} acts like \hyperlink{method.clarify}{\mbox{\isa{clarify}}}, but also does + simplification. Note that if the Simplifier context includes a + splitter for the premises, the subgoal may still be split. \end{description}% \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsubsection{Declaring rules% +\isamarkupsubsection{Structured proof methods% } \isamarkuptrue% % \begin{isamarkuptext}% \begin{matharray}{rcl} - \indexdef{}{command}{print\_claset}\hypertarget{command.print-claset}{\hyperlink{command.print-claset}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}claset}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{attribute}{intro}\hypertarget{attribute.intro}{\hyperlink{attribute.intro}{\mbox{\isa{intro}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{elim}\hypertarget{attribute.elim}{\hyperlink{attribute.elim}{\mbox{\isa{elim}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{dest}\hypertarget{attribute.dest}{\hyperlink{attribute.dest}{\mbox{\isa{dest}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{rule}\hypertarget{attribute.rule}{\hyperlink{attribute.rule}{\mbox{\isa{rule}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{iff}\hypertarget{attribute.iff}{\hyperlink{attribute.iff}{\mbox{\isa{iff}}}} & : & \isa{attribute} \\ + \indexdef{}{method}{rule}\hypertarget{method.rule}{\hyperlink{method.rule}{\mbox{\isa{rule}}}} & : & \isa{method} \\ + \indexdef{}{method}{contradiction}\hypertarget{method.contradiction}{\hyperlink{method.contradiction}{\mbox{\isa{contradiction}}}} & : & \isa{method} \\ + \indexdef{}{method}{intro}\hypertarget{method.intro}{\hyperlink{method.intro}{\mbox{\isa{intro}}}} & : & \isa{method} \\ + \indexdef{}{method}{elim}\hypertarget{method.elim}{\hyperlink{method.elim}{\mbox{\isa{elim}}}} & : & \isa{method} \\ \end{matharray} \begin{railoutput} \rail@begin{3}{} \rail@bar -\rail@term{\hyperlink{attribute.intro}{\mbox{\isa{intro}}}}[] +\rail@term{\hyperlink{method.rule}{\mbox{\isa{rule}}}}[] \rail@nextbar{1} -\rail@term{\hyperlink{attribute.elim}{\mbox{\isa{elim}}}}[] +\rail@term{\hyperlink{method.intro}{\mbox{\isa{intro}}}}[] \rail@nextbar{2} -\rail@term{\hyperlink{attribute.dest}{\mbox{\isa{dest}}}}[] -\rail@endbar -\rail@bar -\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[] -\rail@nextbar{1} -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] +\rail@term{\hyperlink{method.elim}{\mbox{\isa{elim}}}}[] \rail@endbar \rail@bar \rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{attribute.rule}{\mbox{\isa{rule}}}}[] -\rail@term{\isa{del}}[] -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{attribute.iff}{\mbox{\isa{iff}}}}[] -\rail@bar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] -\rail@endbar -\rail@nextbar{2} -\rail@term{\isa{del}}[] +\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] \rail@endbar \rail@end \end{railoutput} @@ -1191,50 +1637,26 @@ \begin{description} - \item \hyperlink{command.print-claset}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}claset}}}} prints the collection of rules - declared to the Classical Reasoner, which is also known as - ``claset'' internally \cite{isabelle-ref}. - - \item \hyperlink{attribute.intro}{\mbox{\isa{intro}}}, \hyperlink{attribute.elim}{\mbox{\isa{elim}}}, and \hyperlink{attribute.dest}{\mbox{\isa{dest}}} - declare introduction, elimination, and destruction rules, - respectively. By default, rules are considered as \emph{unsafe} - (i.e.\ not applied blindly without backtracking), while ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}}'' classifies as \emph{safe}. Rule declarations marked by - ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'' coincide with those of Isabelle/Pure, cf.\ - \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps - of the \hyperlink{method.rule}{\mbox{\isa{rule}}} method). The optional natural number - specifies an explicit weight argument, which is ignored by automated - tools, but determines the search order of single rule steps. + \item \hyperlink{method.rule}{\mbox{\isa{rule}}} as offered by the Classical Reasoner is a + refinement over the Pure one (see \secref{sec:pure-meth-att}). Both + versions work the same, but the classical version observes the + classical rule context in addition to that of Isabelle/Pure. - \item \hyperlink{attribute.rule}{\mbox{\isa{rule}}}~\isa{del} deletes introduction, - elimination, or destruction rules from the context. + Common object logics (HOL, ZF, etc.) declare a rich collection of + classical rules (even if these would qualify as intuitionistic + ones), but only few declarations to the rule context of + Isabelle/Pure (\secref{sec:pure-meth-att}). - \item \hyperlink{attribute.iff}{\mbox{\isa{iff}}} declares logical equivalences to the - Simplifier and the Classical reasoner at the same time. - Non-conditional rules result in a ``safe'' introduction and - elimination pair; conditional ones are considered ``unsafe''. Rules - with negative conclusion are automatically inverted (using \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}{\isaliteral{22}{\isachardoublequote}}}-elimination internally). - - The ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'' version of \hyperlink{attribute.iff}{\mbox{\isa{iff}}} declares rules to - the Isabelle/Pure context only, and omits the Simplifier - declaration. + \item \hyperlink{method.contradiction}{\mbox{\isa{contradiction}}} solves some goal by contradiction, + deriving any result from both \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequote}}} and \isa{A}. Chained + facts, which are guaranteed to participate, may appear in either + order. - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Classical operations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{attribute}{swapped}\hypertarget{attribute.swapped}{\hyperlink{attribute.swapped}{\mbox{\isa{swapped}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{description} - - \item \hyperlink{attribute.swapped}{\mbox{\isa{swapped}}} turns an introduction rule into an - elimination, by resolving with the classical swap principle \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. + \item \hyperlink{method.intro}{\mbox{\isa{intro}}} and \hyperlink{method.elim}{\mbox{\isa{elim}}} repeatedly refine some goal + by intro- or elim-resolution, after having inserted any chained + facts. Exactly the rules given as arguments are taken into account; + this allows fine-tuned decomposition of a proof problem, in contrast + to common automated tools. \end{description}% \end{isamarkuptext}% diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/document/HOL_Specific.tex --- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex Wed Jun 08 10:24:07 2011 +0200 @@ -22,52 +22,383 @@ } \isamarkuptrue% % -\isamarkupsection{Typedef axiomatization \label{sec:hol-typedef}% +\isamarkupsection{Higher-Order Logic% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle/HOL is based on Higher-Order Logic, a polymorphic + version of Church's Simple Theory of Types. HOL can be best + understood as a simply-typed version of classical set theory. The + logic was first implemented in Gordon's HOL system + \cite{mgordon-hol}. It extends Church's original logic + \cite{church40} by explicit type variables (naive polymorphism) and + a sound axiomatization scheme for new types based on subsets of + existing types. + + Andrews's book \cite{andrews86} is a full description of the + original Church-style higher-order logic, with proofs of correctness + and completeness wrt.\ certain set-theoretic interpretations. The + particular extensions of Gordon-style HOL are explained semantically + in two chapters of the 1993 HOL book \cite{pitts93}. + + Experience with HOL over decades has demonstrated that higher-order + logic is widely applicable in many areas of mathematics and computer + science. In a sense, Higher-Order Logic is simpler than First-Order + Logic, because there are fewer restrictions and special cases. Note + that HOL is \emph{weaker} than FOL with axioms for ZF set theory, + which is traditionally considered the standard foundation of regular + mathematics, but for most applications this does not matter. If you + prefer ML to Lisp, you will probably prefer HOL to ZF. + + \medskip The syntax of HOL follows \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-calculus and + functional programming. Function application is curried. To apply + the function \isa{f} of type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{3}}{\isaliteral{22}{\isachardoublequote}}} to the + arguments \isa{a} and \isa{b} in HOL, you simply write \isa{{\isaliteral{22}{\isachardoublequote}}f\ a\ b{\isaliteral{22}{\isachardoublequote}}} (as in ML or Haskell). There is no ``apply'' operator; the + existing application of the Pure \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-calculus is re-used. + Note that in HOL \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{28}{\isacharparenleft}}a{\isaliteral{2C}{\isacharcomma}}\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} means ``\isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{22}{\isachardoublequote}}} applied to + the pair \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}a{\isaliteral{2C}{\isacharcomma}}\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} (which is notation for \isa{{\isaliteral{22}{\isachardoublequote}}Pair\ a\ b{\isaliteral{22}{\isachardoublequote}}}). The latter typically introduces extra formal efforts that can + be avoided by currying functions by default. Explicit tuples are as + infrequent in HOL formalizations as in good ML or Haskell programs. + + \medskip Isabelle/HOL has a distinct feel, compared to other + object-logics like Isabelle/ZF. It identifies object-level types + with meta-level types, taking advantage of the default + type-inference mechanism of Isabelle/Pure. HOL fully identifies + object-level functions with meta-level functions, with native + abstraction and application. + + These identifications allow Isabelle to support HOL particularly + nicely, but they also mean that HOL requires some sophistication + from the user. In particular, an understanding of Hindley-Milner + type-inference with type-classes, which are both used extensively in + the standard libraries and applications. Beginners can set + \hyperlink{attribute.show-types}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}types}}} or even \hyperlink{attribute.show-sorts}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}sorts}}} to get more + explicit information about the result of type-inference.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Inductive and coinductive definitions \label{sec:hol-inductive}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +An \emph{inductive definition} specifies the least predicate + or set \isa{R} closed under given rules: applying a rule to + elements of \isa{R} yields a result within \isa{R}. For + example, a structural operational semantics is an inductive + definition of an evaluation relation. + + Dually, a \emph{coinductive definition} specifies the greatest + predicate or set \isa{R} that is consistent with given rules: + every element of \isa{R} can be seen as arising by applying a rule + to elements of \isa{R}. An important example is using + bisimulation relations to formalise equivalence of processes and + infinite data structures. + + Both inductive and coinductive definitions are based on the + Knaster-Tarski fixed-point theorem for complete lattices. The + collection of introduction rules given by the user determines a + functor on subsets of set-theoretic relations. The required + monotonicity of the recursion scheme is proven as a prerequisite to + the fixed-point definition and the resulting consequences. This + works by pushing inclusion through logical connectives and any other + operator that might be wrapped around recursive occurrences of the + defined relation: there must be a monotonicity theorem of the form + \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C4D3E}{\isasymM}}\ A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{5C3C4D3E}{\isasymM}}\ B{\isaliteral{22}{\isachardoublequote}}}, for each premise \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4D3E}{\isasymM}}\ R\ t{\isaliteral{22}{\isachardoublequote}}} in an + introduction rule. The default rule declarations of Isabelle/HOL + already take care of most common situations. + + \begin{matharray}{rcl} + \indexdef{HOL}{command}{inductive}\hypertarget{command.HOL.inductive}{\hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{HOL}{command}{inductive\_set}\hypertarget{command.HOL.inductive-set}{\hyperlink{command.HOL.inductive-set}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{HOL}{command}{coinductive}\hypertarget{command.HOL.coinductive}{\hyperlink{command.HOL.coinductive}{\mbox{\isa{\isacommand{coinductive}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{HOL}{command}{coinductive\_set}\hypertarget{command.HOL.coinductive-set}{\hyperlink{command.HOL.coinductive-set}{\mbox{\isa{\isacommand{coinductive{\isaliteral{5F}{\isacharunderscore}}set}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{HOL}{attribute}{mono}\hypertarget{attribute.HOL.mono}{\hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}}} & : & \isa{attribute} \\ + \end{matharray} + + \begin{railoutput} +\rail@begin{10}{} +\rail@bar +\rail@term{\hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}}}[] +\rail@nextbar{1} +\rail@term{\hyperlink{command.HOL.inductive-set}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}}}}}[] +\rail@nextbar{2} +\rail@term{\hyperlink{command.HOL.coinductive}{\mbox{\isa{\isacommand{coinductive}}}}}[] +\rail@nextbar{3} +\rail@term{\hyperlink{command.HOL.coinductive-set}{\mbox{\isa{\isacommand{coinductive{\isaliteral{5F}{\isacharunderscore}}set}}}}}[] +\rail@endbar +\rail@bar +\rail@nextbar{1} +\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] +\rail@endbar +\rail@cr{5} +\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] +\rail@bar +\rail@nextbar{6} +\rail@term{\isa{\isakeyword{for}}}[] +\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] +\rail@endbar +\rail@bar +\rail@nextbar{6} +\rail@term{\isa{\isakeyword{where}}}[] +\rail@nont{\isa{clauses}}[] +\rail@endbar +\rail@cr{8} +\rail@bar +\rail@nextbar{9} +\rail@term{\isa{\isakeyword{monos}}}[] +\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] +\rail@endbar +\rail@end +\rail@begin{3}{\isa{clauses}} +\rail@plus +\rail@bar +\rail@nextbar{1} +\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] +\rail@endbar +\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] +\rail@nextplus{2} +\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] +\rail@endplus +\rail@end +\rail@begin{3}{} +\rail@term{\hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}}}[] +\rail@bar +\rail@nextbar{1} +\rail@term{\isa{add}}[] +\rail@nextbar{2} +\rail@term{\isa{del}}[] +\rail@endbar +\rail@end +\end{railoutput} + + + \begin{description} + + \item \hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}} and \hyperlink{command.HOL.coinductive}{\mbox{\isa{\isacommand{coinductive}}}} define (co)inductive predicates from the introduction + rules. + + The propositions given as \isa{{\isaliteral{22}{\isachardoublequote}}clauses{\isaliteral{22}{\isachardoublequote}}} in the \hyperlink{keyword.where}{\mbox{\isa{\isakeyword{where}}}} part are either rules of the usual \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{2F}{\isacharslash}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} format + (with arbitrary nesting), or equalities using \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}}. The + latter specifies extra-logical abbreviations in the sense of + \indexref{}{command}{abbreviation}\hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}}. Introducing abstract syntax + simultaneously with the actual introduction rules is occasionally + useful for complex specifications. + + The optional \hyperlink{keyword.for}{\mbox{\isa{\isakeyword{for}}}} part contains a list of parameters of + the (co)inductive predicates that remain fixed throughout the + definition, in contrast to arguments of the relation that may vary + in each occurrence within the given \isa{{\isaliteral{22}{\isachardoublequote}}clauses{\isaliteral{22}{\isachardoublequote}}}. + + The optional \hyperlink{keyword.monos}{\mbox{\isa{\isakeyword{monos}}}} declaration contains additional + \emph{monotonicity theorems}, which are required for each operator + applied to a recursive set in the introduction rules. + + \item \hyperlink{command.HOL.inductive-set}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}}}} and \hyperlink{command.HOL.coinductive-set}{\mbox{\isa{\isacommand{coinductive{\isaliteral{5F}{\isacharunderscore}}set}}}} are wrappers for to the previous commands for + native HOL predicates. This allows to define (co)inductive sets, + where multiple arguments are simulated via tuples. + + \item \hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}} declares monotonicity rules in the + context. These rule are involved in the automated monotonicity + proof of the above inductive and coinductive definitions. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Derived rules% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +A (co)inductive definition of \isa{R} provides the following + main theorems: + + \begin{description} + + \item \isa{R{\isaliteral{2E}{\isachardot}}intros} is the list of introduction rules as proven + theorems, for the recursive predicates (or sets). The rules are + also available individually, using the names given them in the + theory file; + + \item \isa{R{\isaliteral{2E}{\isachardot}}cases} is the case analysis (or elimination) rule; + + \item \isa{R{\isaliteral{2E}{\isachardot}}induct} or \isa{R{\isaliteral{2E}{\isachardot}}coinduct} is the (co)induction + rule. + + \end{description} + + When several predicates \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} are + defined simultaneously, the list of introduction rules is called + \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5F}{\isacharunderscore}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}intros{\isaliteral{22}{\isachardoublequote}}}, the case analysis rules are + called \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2E}{\isachardot}}cases{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}cases{\isaliteral{22}{\isachardoublequote}}}, and the list + of mutual induction rules is called \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5F}{\isacharunderscore}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}inducts{\isaliteral{22}{\isachardoublequote}}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Monotonicity theorems% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The context maintains a default set of theorems that are used + in monotonicity proofs. New rules can be declared via the + \hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}} attribute. See the main Isabelle/HOL + sources for some examples. The general format of such monotonicity + theorems is as follows: + + \begin{itemize} + + \item Theorems of the form \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C4D3E}{\isasymM}}\ A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{5C3C4D3E}{\isasymM}}\ B{\isaliteral{22}{\isachardoublequote}}}, for proving + monotonicity of inductive definitions whose introduction rules have + premises involving terms such as \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4D3E}{\isasymM}}\ R\ t{\isaliteral{22}{\isachardoublequote}}}. + + \item Monotonicity theorems for logical operators, which are of the + general form \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}. For example, in + the case of the operator \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6F723E}{\isasymor}}{\isaliteral{22}{\isachardoublequote}}}, the corresponding theorem is + \[ + \infer{\isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}} + \] + + \item De Morgan style equations for reasoning about the ``polarity'' + of expressions, e.g. + \[ + \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ P{\isaliteral{22}{\isachardoublequote}}} \qquad\qquad + \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C616E643E}{\isasymand}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q{\isaliteral{22}{\isachardoublequote}}} + \] + + \item Equations for reducing complex operators to more primitive + ones whose monotonicity can easily be proved, e.g. + \[ + \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}} \qquad\qquad + \isa{{\isaliteral{22}{\isachardoublequote}}Ball\ A\ P\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P\ x{\isaliteral{22}{\isachardoublequote}}} + \] + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Examples% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The finite powerset operator can be defined inductively like this:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}\isamarkupfalse% +\ Fin\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ set\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ set\ set{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{for}\ A\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isakeyword{where}\isanewline +\ \ empty{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{7D}{\isacharbraceright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ Fin\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ insert{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C696E3E}{\isasymin}}\ Fin\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ insert\ a\ B\ {\isaliteral{5C3C696E3E}{\isasymin}}\ Fin\ A{\isaliteral{22}{\isachardoublequoteclose}}% +\begin{isamarkuptext}% +The accessible part of a relation is defined as follows:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{inductive}\isamarkupfalse% +\ acc\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isakeyword{for}\ r\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infix}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C707265633E}{\isasymprec}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\isanewline +\isakeyword{where}\ acc{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}y{\isaliteral{2E}{\isachardot}}\ y\ {\isaliteral{5C3C707265633E}{\isasymprec}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ acc\ r\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ acc\ r\ x{\isaliteral{22}{\isachardoublequoteclose}}% +\begin{isamarkuptext}% +Common logical connectives can be easily characterized as +non-recursive inductive definitions with parameters, but without +arguments.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{inductive}\isamarkupfalse% +\ AND\ \isakeyword{for}\ A\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ bool\isanewline +\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ AND\ A\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isanewline +\isacommand{inductive}\isamarkupfalse% +\ OR\ \isakeyword{for}\ A\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ bool\isanewline +\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ OR\ A\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ OR\ A\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isanewline +\isacommand{inductive}\isamarkupfalse% +\ EXISTS\ \isakeyword{for}\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ a\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ EXISTS\ B{\isaliteral{22}{\isachardoublequoteclose}}% +\begin{isamarkuptext}% +Here the \isa{{\isaliteral{22}{\isachardoublequote}}cases{\isaliteral{22}{\isachardoublequote}}} or \isa{{\isaliteral{22}{\isachardoublequote}}induct{\isaliteral{22}{\isachardoublequote}}} rules produced by + the \hyperlink{command.inductive}{\mbox{\isa{\isacommand{inductive}}}} package coincide with the expected + elimination rules for Natural Deduction. Already in the original + article by Gerhard Gentzen \cite{Gentzen:1935} there is a hint that + each connective can be characterized by its introductions, and the + elimination can be constructed systematically.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Recursive functions \label{sec:recursion}% } \isamarkuptrue% % \begin{isamarkuptext}% \begin{matharray}{rcl} - \indexdef{HOL}{command}{typedef}\hypertarget{command.HOL.typedef}{\hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{HOL}{command}{primrec}\hypertarget{command.HOL.primrec}{\hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{HOL}{command}{fun}\hypertarget{command.HOL.fun}{\hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{HOL}{command}{function}\hypertarget{command.HOL.function}{\hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{HOL}{command}{termination}\hypertarget{command.HOL.termination}{\hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ \end{matharray} \begin{railoutput} \rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}}[] +\rail@term{\hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}}}[] +\rail@bar +\rail@nextbar{1} +\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] +\rail@endbar +\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] +\rail@term{\isa{\isakeyword{where}}}[] +\rail@nont{\isa{equations}}[] +\rail@end +\rail@begin{4}{} +\rail@bar +\rail@term{\hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}}}[] +\rail@nextbar{1} +\rail@term{\hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}}[] +\rail@endbar +\rail@bar +\rail@nextbar{1} +\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] +\rail@endbar \rail@bar \rail@nextbar{1} -\rail@nont{\isa{altname}}[] +\rail@nont{\isa{functionopts}}[] \rail@endbar -\rail@nont{\isa{abstype}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\isa{repset}}[] +\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] +\rail@cr{3} +\rail@term{\isa{\isakeyword{where}}}[] +\rail@nont{\isa{equations}}[] \rail@end -\rail@begin{3}{\isa{altname}} +\rail@begin{3}{\isa{equations}} +\rail@plus +\rail@bar +\rail@nextbar{1} +\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] +\rail@endbar +\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] +\rail@nextplus{2} +\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] +\rail@endplus +\rail@end +\rail@begin{3}{\isa{functionopts}} \rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] +\rail@plus \rail@bar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] +\rail@term{\isa{sequential}}[] \rail@nextbar{1} -\rail@term{\isa{\isakeyword{open}}}[] -\rail@nextbar{2} -\rail@term{\isa{\isakeyword{open}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] +\rail@term{\isa{domintros}}[] \rail@endbar +\rail@nextplus{2} +\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] +\rail@endplus \rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] \rail@end -\rail@begin{2}{\isa{abstype}} -\rail@nont{\hyperlink{syntax.typespec-sorts}{\mbox{\isa{typespec{\isaliteral{5F}{\isacharunderscore}}sorts}}}}[] +\rail@begin{2}{} +\rail@term{\hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}}[] \rail@bar \rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\isa{repset}} \rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{morphisms}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] \rail@endbar \rail@end \end{railoutput} @@ -75,60 +406,540 @@ \begin{description} - \item \hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{3D}{\isacharequal}}\ A{\isaliteral{22}{\isachardoublequote}}} - axiomatizes a Gordon/HOL-style type definition in the background - theory of the current context, depending on a non-emptiness result - of the set \isa{A} (which needs to be proven interactively). + \item \hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}} defines primitive recursive + functions over datatypes (see also \indexref{HOL}{command}{datatype}\hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}} and + \indexref{HOL}{command}{rep\_datatype}\hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}}). The given \isa{equations} + specify reduction rules that are produced by instantiating the + generic combinator for primitive recursion that is available for + each datatype. + + Each equation needs to be of the form: + + \begin{isabelle}% +{\isaliteral{22}{\isachardoublequote}}f\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{28}{\isacharparenleft}}C\ y\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ y\isaliteral{5C3C5E7375623E}{}\isactrlsub k{\isaliteral{29}{\isacharparenright}}\ z\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ z\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3D}{\isacharequal}}\ rhs{\isaliteral{22}{\isachardoublequote}}% +\end{isabelle} - The raw type may not depend on parameters or assumptions of the - context --- this is logically impossible in Isabelle/HOL --- but the - non-emptiness property can be local, potentially resulting in - multiple interpretations in target contexts. Thus the established - bijection between the representing set \isa{A} and the new type - \isa{t} may semantically depend on local assumptions. + such that \isa{C} is a datatype constructor, \isa{rhs} contains + only the free variables on the left-hand side (or from the context), + and all recursive occurrences of \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{22}{\isachardoublequote}}} in \isa{{\isaliteral{22}{\isachardoublequote}}rhs{\isaliteral{22}{\isachardoublequote}}} are of + the form \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ y\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} for some \isa{i}. At most one + reduction rule for each constructor can be given. The order does + not matter. For missing constructors, the function is defined to + return a default value, but this equation is made difficult to + access for users. + + The reduction rules are declared as \hyperlink{attribute.simp}{\mbox{\isa{simp}}} by default, + which enables standard proof methods like \hyperlink{method.simp}{\mbox{\isa{simp}}} and + \hyperlink{method.auto}{\mbox{\isa{auto}}} to normalize expressions of \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{22}{\isachardoublequote}}} applied to + datatype constructions, by simulating symbolic computation via + rewriting. + + \item \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} defines functions by general + wellfounded recursion. A detailed description with examples can be + found in \cite{isabelle-function}. The function is specified by a + set of (possibly conditional) recursive equations with arbitrary + pattern matching. The command generates proof obligations for the + completeness and the compatibility of patterns. + + The defined function is considered partial, and the resulting + simplification rules (named \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}psimps{\isaliteral{22}{\isachardoublequote}}}) and induction rule + (named \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}pinduct{\isaliteral{22}{\isachardoublequote}}}) are guarded by a generated domain + predicate \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{5F}{\isacharunderscore}}dom{\isaliteral{22}{\isachardoublequote}}}. The \hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}} + command can then be used to establish that the function is total. - By default, \hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}} defines both a type \isa{t} - and a set (term constant) of the same name, unless an alternative - base name is given in parentheses, or the ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}open{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' - declaration is used to suppress a separate constant definition - altogether. The injection from type to set is called \isa{Rep{\isaliteral{5F}{\isacharunderscore}}t}, - its inverse \isa{Abs{\isaliteral{5F}{\isacharunderscore}}t} --- this may be changed via an explicit - \hyperlink{keyword.HOL.morphisms}{\mbox{\isa{\isakeyword{morphisms}}}} declaration. + \item \hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}} is a shorthand notation for ``\hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}sequential{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, followed by automated + proof attempts regarding pattern matching and termination. See + \cite{isabelle-function} for further details. + + \item \hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}~\isa{f} commences a + termination proof for the previously defined function \isa{f}. If + this is omitted, the command refers to the most recent function + definition. After the proof is closed, the recursive equations and + the induction principle is established. + + \end{description} + + Recursive definitions introduced by the \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} + command accommodate reasoning by induction (cf.\ \hyperlink{method.induct}{\mbox{\isa{induct}}}): + rule \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}induct{\isaliteral{22}{\isachardoublequote}}} refers to a specific induction rule, with + parameters named according to the user-specified equations. Cases + are numbered starting from 1. For \hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}}, the + induction principle coincides with structural recursion on the + datatype where the recursion is carried out. - Theorems \isa{Rep{\isaliteral{5F}{\isacharunderscore}}t}, \isa{Rep{\isaliteral{5F}{\isacharunderscore}}t{\isaliteral{5F}{\isacharunderscore}}inverse}, and \isa{Abs{\isaliteral{5F}{\isacharunderscore}}t{\isaliteral{5F}{\isacharunderscore}}inverse} provide the most basic characterization as a - corresponding injection/surjection pair (in both directions). Rules - \isa{Rep{\isaliteral{5F}{\isacharunderscore}}t{\isaliteral{5F}{\isacharunderscore}}inject} and \isa{Abs{\isaliteral{5F}{\isacharunderscore}}t{\isaliteral{5F}{\isacharunderscore}}inject} provide a slightly - more convenient view on the injectivity part, suitable for automated - proof tools (e.g.\ in \hyperlink{attribute.simp}{\mbox{\isa{simp}}} or \hyperlink{attribute.iff}{\mbox{\isa{iff}}} - declarations). Rules \isa{Rep{\isaliteral{5F}{\isacharunderscore}}t{\isaliteral{5F}{\isacharunderscore}}cases}/\isa{Rep{\isaliteral{5F}{\isacharunderscore}}t{\isaliteral{5F}{\isacharunderscore}}induct}, and - \isa{Abs{\isaliteral{5F}{\isacharunderscore}}t{\isaliteral{5F}{\isacharunderscore}}cases}/\isa{Abs{\isaliteral{5F}{\isacharunderscore}}t{\isaliteral{5F}{\isacharunderscore}}induct} provide alternative views - on surjectivity; these are already declared as set or type rules for - the generic \hyperlink{method.cases}{\mbox{\isa{cases}}} and \hyperlink{method.induct}{\mbox{\isa{induct}}} methods. + The equations provided by these packages may be referred later as + theorem list \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}simps{\isaliteral{22}{\isachardoublequote}}}, where \isa{f} is the (collective) + name of the functions defined. Individual equations may be named + explicitly as well. + + The \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} command accepts the following + options. + + \begin{description} - An alternative name for the set definition (and other derived - entities) may be specified in parentheses; the default is to use - \isa{t} as indicated before. + \item \isa{sequential} enables a preprocessor which disambiguates + overlapping patterns by making them mutually disjoint. Earlier + equations take precedence over later ones. This allows to give the + specification in a format very similar to functional programming. + Note that the resulting simplification and induction rules + correspond to the transformed specification, not the one given + originally. This usually means that each equation given by the user + may result in several theorems. Also note that this automatic + transformation only works for ML-style datatype patterns. + + \item \isa{domintros} enables the automated generation of + introduction rules for the domain predicate. While mostly not + needed, they can be helpful in some proofs about partial functions. \end{description}% \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsection{Adhoc tuples% +\isamarkupsubsubsection{Example: evaluation of expressions% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Subsequently, we define mutual datatypes for arithmetic and + boolean expressions, and use \hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}} for evaluation + functions that follow the same recursive structure.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ {\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{3D}{\isacharequal}}\isanewline +\ \ \ \ IF\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ {\isaliteral{7C}{\isacharbar}}\ Sum\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ {\isaliteral{7C}{\isacharbar}}\ Diff\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ {\isaliteral{7C}{\isacharbar}}\ Var\ {\isaliteral{27}{\isacharprime}}a\isanewline +\ \ {\isaliteral{7C}{\isacharbar}}\ Num\ nat\isanewline +\isakeyword{and}\ {\isaliteral{27}{\isacharprime}}a\ bexp\ {\isaliteral{3D}{\isacharequal}}\isanewline +\ \ \ \ Less\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ {\isaliteral{7C}{\isacharbar}}\ And\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ {\isaliteral{7C}{\isacharbar}}\ Neg\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}% +\begin{isamarkuptext}% +\medskip Evaluation of arithmetic and boolean expressions% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\ evala\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isakeyword{and}\ evalb\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ bexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isakeyword{where}\isanewline +\ \ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}IF\ b\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}if\ evalb\ env\ b\ then\ evala\ env\ a{\isadigit{1}}\ else\ evala\ env\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}Sum\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evala\ env\ a{\isadigit{1}}\ {\isaliteral{2B}{\isacharplus}}\ evala\ env\ a{\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}Diff\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evala\ env\ a{\isadigit{1}}\ {\isaliteral{2D}{\isacharminus}}\ evala\ env\ a{\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}Var\ v{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ env\ v{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}Num\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}Less\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}evala\ env\ a{\isadigit{1}}\ {\isaliteral{3C}{\isacharless}}\ evala\ env\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}And\ b{\isadigit{1}}\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}evalb\ env\ b{\isadigit{1}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ evalb\ env\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}Neg\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ evalb\ env\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% +\begin{isamarkuptext}% +Since the value of an expression depends on the value of its + variables, the functions \isa{evala} and \isa{evalb} take an + additional parameter, an \emph{environment} that maps variables to + their values. + + \medskip Substitution on expressions can be defined similarly. The + mapping \isa{f} of type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequote}}} given as a + parameter is lifted canonically on the types \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequote}}} and + \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequote}}}, respectively.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\ substa\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ aexp{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isakeyword{and}\ substb\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ aexp{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ bexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isakeyword{where}\isanewline +\ \ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}IF\ b\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ IF\ {\isaliteral{28}{\isacharparenleft}}substb\ f\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}Sum\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Sum\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}Diff\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Diff\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}Var\ v{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ f\ v{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}Num\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Num\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substb\ f\ {\isaliteral{28}{\isacharparenleft}}Less\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Less\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substb\ f\ {\isaliteral{28}{\isacharparenleft}}And\ b{\isadigit{1}}\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ And\ {\isaliteral{28}{\isacharparenleft}}substb\ f\ b{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substb\ f\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substb\ f\ {\isaliteral{28}{\isacharparenleft}}Neg\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Neg\ {\isaliteral{28}{\isacharparenleft}}substb\ f\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% +\begin{isamarkuptext}% +In textbooks about semantics one often finds substitution + theorems, which express the relationship between substitution and + evaluation. For \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequote}}}, we can prove + such a theorem by mutual induction, followed by simplification.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ subst{\isaliteral{5F}{\isacharunderscore}}one{\isaliteral{3A}{\isacharcolon}}\isanewline +\ \ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}substa\ {\isaliteral{28}{\isacharparenleft}}Var\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evala\ {\isaliteral{28}{\isacharparenleft}}env\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ evala\ env\ a{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}substb\ {\isaliteral{28}{\isacharparenleft}}Var\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evalb\ {\isaliteral{28}{\isacharparenleft}}env\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ evala\ env\ a{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ a\ \isakeyword{and}\ b{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ subst{\isaliteral{5F}{\isacharunderscore}}all{\isaliteral{3A}{\isacharcolon}}\isanewline +\ \ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}substa\ s\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evala\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ evala\ env\ {\isaliteral{28}{\isacharparenleft}}s\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}substb\ s\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evalb\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ evala\ env\ {\isaliteral{28}{\isacharparenleft}}s\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ a\ \isakeyword{and}\ b{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsubsection{Example: a substitution function for terms% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Functions on datatypes with nested recursion are also defined + by mutual primitive recursion.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{22}{\isachardoublequoteopen}}term{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{3D}{\isacharequal}}\ Var\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{7C}{\isacharbar}}\ App\ {\isaliteral{27}{\isacharprime}}b\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ list{\isaliteral{22}{\isachardoublequoteclose}}% +\begin{isamarkuptext}% +A substitution function on type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{22}{\isachardoublequote}}} can be + defined as follows, by working simultaneously on \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ list{\isaliteral{22}{\isachardoublequote}}}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{primrec}\isamarkupfalse% +\ subst{\isaliteral{5F}{\isacharunderscore}}term\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline +\ \ subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ list{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isakeyword{where}\isanewline +\ \ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f\ {\isaliteral{28}{\isacharparenleft}}Var\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ f\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f\ {\isaliteral{28}{\isacharparenleft}}App\ b\ ts{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ App\ b\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f\ ts{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f\ {\isaliteral{28}{\isacharparenleft}}t\ {\isaliteral{23}{\isacharhash}}\ ts{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ subst{\isaliteral{5F}{\isacharunderscore}}term\ f\ t\ {\isaliteral{23}{\isacharhash}}\ subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f\ ts{\isaliteral{22}{\isachardoublequoteclose}}% +\begin{isamarkuptext}% +The recursion scheme follows the structure of the unfolded + definition of type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{22}{\isachardoublequote}}}. To prove properties of this + substitution function, mutual induction is needed:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f{\isadigit{1}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ f{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{3D}{\isacharequal}}\ subst{\isaliteral{5F}{\isacharunderscore}}term\ f{\isadigit{1}}\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f{\isadigit{2}}\ t{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline +\ \ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f{\isadigit{1}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ f{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ ts\ {\isaliteral{3D}{\isacharequal}}\ subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f{\isadigit{1}}\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f{\isadigit{2}}\ ts{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ t\ \isakeyword{and}\ ts{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsubsection{Example: a map function for infinitely branching trees% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Defining functions on infinitely branching datatypes by + primitive recursion is just as easy.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ {\isaliteral{27}{\isacharprime}}a\ tree\ {\isaliteral{3D}{\isacharequal}}\ Atom\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{7C}{\isacharbar}}\ Branch\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ tree{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isanewline +\isacommand{primrec}\isamarkupfalse% +\ map{\isaliteral{5F}{\isacharunderscore}}tree\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ tree\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ tree{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isakeyword{where}\isanewline +\ \ {\isaliteral{22}{\isachardoublequoteopen}}map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ {\isaliteral{28}{\isacharparenleft}}Atom\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Atom\ {\isaliteral{28}{\isacharparenleft}}f\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ {\isaliteral{28}{\isacharparenleft}}Branch\ ts{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Branch\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ {\isaliteral{28}{\isacharparenleft}}ts\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% +\begin{isamarkuptext}% +Note that all occurrences of functions such as \isa{ts} + above must be applied to an argument. In particular, \isa{{\isaliteral{22}{\isachardoublequote}}map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ ts{\isaliteral{22}{\isachardoublequote}}} is not allowed here.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +Here is a simple composition lemma for \isa{map{\isaliteral{5F}{\isacharunderscore}}tree}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}map{\isaliteral{5F}{\isacharunderscore}}tree\ g\ {\isaliteral{28}{\isacharparenleft}}map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ t{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ map{\isaliteral{5F}{\isacharunderscore}}tree\ {\isaliteral{28}{\isacharparenleft}}g\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ f{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ t{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Proof methods related to recursive definitions% } \isamarkuptrue% % \begin{isamarkuptext}% \begin{matharray}{rcl} - \indexdef{HOL}{attribute}{split\_format}\hypertarget{attribute.HOL.split-format}{\hyperlink{attribute.HOL.split-format}{\mbox{\isa{split{\isaliteral{5F}{\isacharunderscore}}format}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{attribute} \\ + \indexdef{HOL}{method}{pat\_completeness}\hypertarget{method.HOL.pat-completeness}{\hyperlink{method.HOL.pat-completeness}{\mbox{\isa{pat{\isaliteral{5F}{\isacharunderscore}}completeness}}}} & : & \isa{method} \\ + \indexdef{HOL}{method}{relation}\hypertarget{method.HOL.relation}{\hyperlink{method.HOL.relation}{\mbox{\isa{relation}}}} & : & \isa{method} \\ + \indexdef{HOL}{method}{lexicographic\_order}\hypertarget{method.HOL.lexicographic-order}{\hyperlink{method.HOL.lexicographic-order}{\mbox{\isa{lexicographic{\isaliteral{5F}{\isacharunderscore}}order}}}} & : & \isa{method} \\ + \indexdef{HOL}{method}{size\_change}\hypertarget{method.HOL.size-change}{\hyperlink{method.HOL.size-change}{\mbox{\isa{size{\isaliteral{5F}{\isacharunderscore}}change}}}} & : & \isa{method} \\ + \end{matharray} + + \begin{railoutput} +\rail@begin{1}{} +\rail@term{\hyperlink{method.HOL.relation}{\mbox{\isa{relation}}}}[] +\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] +\rail@end +\rail@begin{2}{} +\rail@term{\hyperlink{method.HOL.lexicographic-order}{\mbox{\isa{lexicographic{\isaliteral{5F}{\isacharunderscore}}order}}}}[] +\rail@plus +\rail@nextplus{1} +\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] +\rail@endplus +\rail@end +\rail@begin{2}{} +\rail@term{\hyperlink{method.HOL.size-change}{\mbox{\isa{size{\isaliteral{5F}{\isacharunderscore}}change}}}}[] +\rail@nont{\isa{orders}}[] +\rail@plus +\rail@nextplus{1} +\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] +\rail@endplus +\rail@end +\rail@begin{4}{\isa{orders}} +\rail@plus +\rail@nextplus{1} +\rail@bar +\rail@term{\isa{max}}[] +\rail@nextbar{2} +\rail@term{\isa{min}}[] +\rail@nextbar{3} +\rail@term{\isa{ms}}[] +\rail@endbar +\rail@endplus +\rail@end +\end{railoutput} + + + \begin{description} + + \item \hyperlink{method.HOL.pat-completeness}{\mbox{\isa{pat{\isaliteral{5F}{\isacharunderscore}}completeness}}} is a specialized method to + solve goals regarding the completeness of pattern matching, as + required by the \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} package (cf.\ + \cite{isabelle-function}). + + \item \hyperlink{method.HOL.relation}{\mbox{\isa{relation}}}~\isa{R} introduces a termination + proof using the relation \isa{R}. The resulting proof state will + contain goals expressing that \isa{R} is wellfounded, and that the + arguments of recursive calls decrease with respect to \isa{R}. + Usually, this method is used as the initial proof step of manual + termination proofs. + + \item \hyperlink{method.HOL.lexicographic-order}{\mbox{\isa{lexicographic{\isaliteral{5F}{\isacharunderscore}}order}}} attempts a fully + automated termination proof by searching for a lexicographic + combination of size measures on the arguments of the function. The + method accepts the same arguments as the \hyperlink{method.auto}{\mbox{\isa{auto}}} method, + which it uses internally to prove local descents. The \hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}} modifiers are accepted (as for \hyperlink{method.auto}{\mbox{\isa{auto}}}). + + In case of failure, extensive information is printed, which can help + to analyse the situation (cf.\ \cite{isabelle-function}). + + \item \hyperlink{method.HOL.size-change}{\mbox{\isa{size{\isaliteral{5F}{\isacharunderscore}}change}}} also works on termination goals, + using a variation of the size-change principle, together with a + graph decomposition technique (see \cite{krauss_phd} for details). + Three kinds of orders are used internally: \isa{max}, \isa{min}, + and \isa{ms} (multiset), which is only available when the theory + \isa{Multiset} is loaded. When no order kinds are given, they are + tried in order. The search for a termination proof uses SAT solving + internally. + + For local descent proofs, the \hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}} modifiers are + accepted (as for \hyperlink{method.auto}{\mbox{\isa{auto}}}). + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Functions with explicit partiality% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{HOL}{command}{partial\_function}\hypertarget{command.HOL.partial-function}{\hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isaliteral{5F}{\isacharunderscore}}function}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{HOL}{attribute}{partial\_function\_mono}\hypertarget{attribute.HOL.partial-function-mono}{\hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}mono}}}} & : & \isa{attribute} \\ \end{matharray} \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.HOL.split-format}{\mbox{\isa{split{\isaliteral{5F}{\isacharunderscore}}format}}}}[] +\rail@begin{5}{} +\rail@term{\hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isaliteral{5F}{\isacharunderscore}}function}}}}}[] +\rail@bar +\rail@nextbar{1} +\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] +\rail@endbar +\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] +\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] +\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] +\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] +\rail@cr{3} +\rail@term{\isa{\isakeyword{where}}}[] +\rail@bar +\rail@nextbar{4} +\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] +\rail@endbar +\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] +\rail@end +\end{railoutput} + + + \begin{description} + + \item \hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isaliteral{5F}{\isacharunderscore}}function}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}mode{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} defines + recursive functions based on fixpoints in complete partial + orders. No termination proof is required from the user or + constructed internally. Instead, the possibility of non-termination + is modelled explicitly in the result type, which contains an + explicit bottom element. + + Pattern matching and mutual recursion are currently not supported. + Thus, the specification consists of a single function described by a + single recursive equation. + + There are no fixed syntactic restrictions on the body of the + function, but the induced functional must be provably monotonic + wrt.\ the underlying order. The monotonicitity proof is performed + internally, and the definition is rejected when it fails. The proof + can be influenced by declaring hints using the + \hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}mono}}} attribute. + + The mandatory \isa{mode} argument specifies the mode of operation + of the command, which directly corresponds to a complete partial + order on the result type. By default, the following modes are + defined: + + \begin{description} + \item \isa{option} defines functions that map into the \isa{option} type. Here, the value \isa{None} is used to model a + non-terminating computation. Monotonicity requires that if \isa{None} is returned by a recursive call, then the overall result + must also be \isa{None}. This is best achieved through the use of + the monadic operator \isa{{\isaliteral{22}{\isachardoublequote}}Option{\isaliteral{2E}{\isachardot}}bind{\isaliteral{22}{\isachardoublequote}}}. + + \item \isa{tailrec} defines functions with an arbitrary result + type and uses the slightly degenerated partial order where \isa{{\isaliteral{22}{\isachardoublequote}}undefined{\isaliteral{22}{\isachardoublequote}}} is the bottom element. Now, monotonicity requires that + if \isa{undefined} is returned by a recursive call, then the + overall result must also be \isa{undefined}. In practice, this is + only satisfied when each recursive call is a tail call, whose result + is directly returned. Thus, this mode of operation allows the + definition of arbitrary tail-recursive functions. + \end{description} + + Experienced users may define new modes by instantiating the locale + \isa{{\isaliteral{22}{\isachardoublequote}}partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}definitions{\isaliteral{22}{\isachardoublequote}}} appropriately. + + \item \hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}mono}}} declares rules for + use in the internal monononicity proofs of partial function + definitions. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Old-style recursive function definitions (TFL)% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The old TFL commands \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} and \hyperlink{command.HOL.recdef-tc}{\mbox{\isa{\isacommand{recdef{\isaliteral{5F}{\isacharunderscore}}tc}}}} for defining recursive are mostly obsolete; \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} or \hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}} should be used instead. + + \begin{matharray}{rcl} + \indexdef{HOL}{command}{recdef}\hypertarget{command.HOL.recdef}{\hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{HOL}{command}{recdef\_tc}\hypertarget{command.HOL.recdef-tc}{\hyperlink{command.HOL.recdef-tc}{\mbox{\isa{\isacommand{recdef{\isaliteral{5F}{\isacharunderscore}}tc}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ + \end{matharray} + + \begin{railoutput} +\rail@begin{5}{} +\rail@term{\hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}}}[] \rail@bar \rail@nextbar{1} \rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{complete}}[] +\rail@term{\isa{\isakeyword{permissive}}}[] +\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] +\rail@endbar +\rail@cr{3} +\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] +\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] +\rail@plus +\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] +\rail@nextplus{4} +\rail@endplus +\rail@bar +\rail@nextbar{4} +\rail@nont{\isa{hints}}[] +\rail@endbar +\rail@end +\rail@begin{2}{} +\rail@nont{\isa{recdeftc}}[] +\rail@bar +\rail@nextbar{1} +\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] +\rail@endbar +\rail@nont{\isa{tc}}[] +\rail@end +\rail@begin{2}{\isa{hints}} +\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] +\rail@term{\isa{\isakeyword{hints}}}[] +\rail@plus +\rail@nextplus{1} +\rail@cnont{\isa{recdefmod}}[] +\rail@endplus +\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] +\rail@end +\rail@begin{4}{\isa{recdefmod}} +\rail@bar +\rail@bar +\rail@term{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}}[] +\rail@nextbar{1} +\rail@term{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}}[] +\rail@nextbar{2} +\rail@term{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf}}[] +\rail@endbar +\rail@bar +\rail@nextbar{1} +\rail@term{\isa{add}}[] +\rail@nextbar{2} +\rail@term{\isa{del}}[] +\rail@endbar +\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] +\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] +\rail@nextbar{3} +\rail@nont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] +\rail@endbar +\rail@end +\rail@begin{2}{\isa{tc}} +\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] +\rail@bar +\rail@nextbar{1} +\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] +\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] \rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] \rail@endbar \rail@end @@ -137,17 +948,235 @@ \begin{description} - \item \hyperlink{attribute.HOL.split-format}{\mbox{\isa{split{\isaliteral{5F}{\isacharunderscore}}format}}}\ \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}complete{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} causes - arguments in function applications to be represented canonically - according to their tuple type structure. + \item \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} defines general well-founded + recursive functions (using the TFL package), see also + \cite{isabelle-HOL}. The ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}permissive{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' option tells + TFL to recover from failed proof attempts, returning unfinished + results. The \isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}, \isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}, and \isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf} hints refer to auxiliary rules to be used in the internal + automated proof process of TFL. Additional \hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}} + declarations may be given to tune the context of the Simplifier + (cf.\ \secref{sec:simplifier}) and Classical reasoner (cf.\ + \secref{sec:classical}). + + \item \hyperlink{command.HOL.recdef-tc}{\mbox{\isa{\isacommand{recdef{\isaliteral{5F}{\isacharunderscore}}tc}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{28}{\isacharparenleft}}i{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} recommences the + proof for leftover termination condition number \isa{i} (default + 1) as generated by a \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} definition of + constant \isa{c}. + + Note that in most cases, \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} is able to finish + its internal proofs without manual intervention. + + \end{description} + + \medskip Hints for \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} may be also declared + globally, using the following attributes. - Note that this operation tends to invent funny names for new local - parameters introduced. + \begin{matharray}{rcl} + \indexdef{HOL}{attribute}{recdef\_simp}\hypertarget{attribute.HOL.recdef-simp}{\hyperlink{attribute.HOL.recdef-simp}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}}}} & : & \isa{attribute} \\ + \indexdef{HOL}{attribute}{recdef\_cong}\hypertarget{attribute.HOL.recdef-cong}{\hyperlink{attribute.HOL.recdef-cong}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}}}} & : & \isa{attribute} \\ + \indexdef{HOL}{attribute}{recdef\_wf}\hypertarget{attribute.HOL.recdef-wf}{\hyperlink{attribute.HOL.recdef-wf}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf}}}} & : & \isa{attribute} \\ + \end{matharray} - \end{description}% + \begin{railoutput} +\rail@begin{3}{} +\rail@bar +\rail@term{\hyperlink{attribute.HOL.recdef-simp}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}}}}[] +\rail@nextbar{1} +\rail@term{\hyperlink{attribute.HOL.recdef-cong}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}}}}[] +\rail@nextbar{2} +\rail@term{\hyperlink{attribute.HOL.recdef-wf}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf}}}}[] +\rail@endbar +\rail@bar +\rail@nextbar{1} +\rail@term{\isa{add}}[] +\rail@nextbar{2} +\rail@term{\isa{del}}[] +\rail@endbar +\rail@end +\end{railoutput}% \end{isamarkuptext}% \isamarkuptrue% % +\isamarkupsection{Datatypes \label{sec:hol-datatype}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{matharray}{rcl} + \indexdef{HOL}{command}{datatype}\hypertarget{command.HOL.datatype}{\hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{HOL}{command}{rep\_datatype}\hypertarget{command.HOL.rep-datatype}{\hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ + \end{matharray} + + \begin{railoutput} +\rail@begin{2}{} +\rail@term{\hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}}}[] +\rail@plus +\rail@nont{\isa{spec}}[] +\rail@nextplus{1} +\rail@cterm{\isa{\isakeyword{and}}}[] +\rail@endplus +\rail@end +\rail@begin{3}{} +\rail@term{\hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}}}[] +\rail@bar +\rail@nextbar{1} +\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] +\rail@plus +\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] +\rail@nextplus{2} +\rail@endplus +\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] +\rail@endbar +\rail@plus +\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] +\rail@nextplus{1} +\rail@endplus +\rail@end +\rail@begin{2}{\isa{spec}} +\rail@bar +\rail@nextbar{1} +\rail@nont{\hyperlink{syntax.parname}{\mbox{\isa{parname}}}}[] +\rail@endbar +\rail@nont{\hyperlink{syntax.typespec}{\mbox{\isa{typespec}}}}[] +\rail@bar +\rail@nextbar{1} +\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] +\rail@endbar +\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] +\rail@plus +\rail@nont{\isa{cons}}[] +\rail@nextplus{1} +\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] +\rail@endplus +\rail@end +\rail@begin{2}{\isa{cons}} +\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] +\rail@plus +\rail@nextplus{1} +\rail@cnont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] +\rail@endplus +\rail@bar +\rail@nextbar{1} +\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] +\rail@endbar +\rail@end +\end{railoutput} + + + \begin{description} + + \item \hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}} defines inductive datatypes in + HOL. + + \item \hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}} represents existing types as + datatypes. + + For foundational reasons, some basic types such as \isa{nat}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2B}{\isacharplus}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{22}{\isachardoublequote}}}, \isa{bool} and \isa{unit} are + introduced by more primitive means using \indexref{}{command}{typedef}\hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}}. To + recover the rich infrastructure of \hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}} (e.g.\ rules + for \hyperlink{method.cases}{\mbox{\isa{cases}}} and \hyperlink{method.induct}{\mbox{\isa{induct}}} and the primitive recursion + combinators), such types may be represented as actual datatypes + later. This is done by specifying the constructors of the desired + type, and giving a proof of the induction rule, distinctness and + injectivity of constructors. + + For example, see \verb|~~/src/HOL/Sum_Type.thy| for the + representation of the primitive sum type as fully-featured datatype. + + \end{description} + + The generated rules for \hyperlink{method.induct}{\mbox{\isa{induct}}} and \hyperlink{method.cases}{\mbox{\isa{cases}}} provide + case names according to the given constructors, while parameters are + named after the types (see also \secref{sec:cases-induct}). + + See \cite{isabelle-HOL} for more details on datatypes, but beware of + the old-style theory syntax being used there! Apart from proper + proof methods for case-analysis and induction, there are also + emulations of ML tactics \hyperlink{method.HOL.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}} and \hyperlink{method.HOL.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}} available, see \secref{sec:hol-induct-tac}; these admit + to refer directly to the internal structure of subgoals (including + internally bound parameters).% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Examples% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We define a type of finite sequences, with slightly different + names than the existing \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ list{\isaliteral{22}{\isachardoublequote}}} that is already in \hyperlink{theory.Main}{\mbox{\isa{Main}}}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ {\isaliteral{27}{\isacharprime}}a\ seq\ {\isaliteral{3D}{\isacharequal}}\ Empty\ {\isaliteral{7C}{\isacharbar}}\ Seq\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ seq{\isaliteral{22}{\isachardoublequoteclose}}% +\begin{isamarkuptext}% +We can now prove some simple lemma by structural induction:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}Seq\ x\ xs\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ xs\ arbitrary{\isaliteral{3A}{\isacharcolon}}\ x{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \isacommand{case}\isamarkupfalse% +\ Empty% +\begin{isamarkuptxt}% +This case can be proved using the simplifier: the freeness + properties of the datatype are already declared as \hyperlink{attribute.simp}{\mbox{\isa{simp}}} rules.% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}Seq\ x\ Empty\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Empty{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{case}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}Seq\ y\ ys{\isaliteral{29}{\isacharparenright}}% +\begin{isamarkuptxt}% +The step case is proved similarly.% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}Seq\ x\ {\isaliteral{28}{\isacharparenleft}}Seq\ y\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Seq\ y\ ys{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \isacommand{using}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}Seq\ y\ ys\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ ys{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Here is a more succinct version of the same proof:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}Seq\ x\ xs\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ xs\ arbitrary{\isaliteral{3A}{\isacharcolon}}\ x{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% \isamarkupsection{Records \label{sec:hol-record}% } \isamarkuptrue% @@ -400,67 +1429,115 @@ \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsection{Datatypes \label{sec:hol-datatype}% +\isamarkupsubsubsection{Examples% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +See \verb|~~/src/HOL/ex/Records.thy|, for example.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Adhoc tuples% } \isamarkuptrue% % \begin{isamarkuptext}% \begin{matharray}{rcl} - \indexdef{HOL}{command}{datatype}\hypertarget{command.HOL.datatype}{\hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{rep\_datatype}\hypertarget{command.HOL.rep-datatype}{\hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ + \indexdef{HOL}{attribute}{split\_format}\hypertarget{attribute.HOL.split-format}{\hyperlink{attribute.HOL.split-format}{\mbox{\isa{split{\isaliteral{5F}{\isacharunderscore}}format}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{attribute} \\ \end{matharray} \begin{railoutput} \rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}}}[] -\rail@plus -\rail@nont{\isa{spec}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}}}[] +\rail@term{\hyperlink{attribute.HOL.split-format}{\mbox{\isa{split{\isaliteral{5F}{\isacharunderscore}}format}}}}[] \rail@bar \rail@nextbar{1} \rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{2} -\rail@endplus +\rail@term{\isa{complete}}[] \rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] \rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextplus{1} -\rail@endplus \rail@end -\rail@begin{2}{\isa{spec}} +\end{railoutput} + + + \begin{description} + + \item \hyperlink{attribute.HOL.split-format}{\mbox{\isa{split{\isaliteral{5F}{\isacharunderscore}}format}}}\ \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}complete{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} causes + arguments in function applications to be represented canonically + according to their tuple type structure. + + Note that this operation tends to invent funny names for new local + parameters introduced. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Typedef axiomatization \label{sec:hol-typedef}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +A Gordon/HOL-style type definition is a certain axiom scheme + that identifies a new type with a subset of an existing type. More + precisely, the new type is defined by exhibiting an existing type + \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}}, a set \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\ set{\isaliteral{22}{\isachardoublequote}}}, and a theorem that proves + \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequote}}}. Thus \isa{A} is a non-empty subset of \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}}, and the new type denotes this subset. New functions are + postulated that establish an isomorphism between the new type and + the subset. In general, the type \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} may involve type + variables \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} which means that the type definition + produces a type constructor \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} depending on + those type arguments. + + The axiomatization can be considered a ``definition'' in the sense + of the particular set-theoretic interpretation of HOL + \cite{pitts93}, where the universe of types is required to be + downwards-closed wrt.\ arbitrary non-empty subsets. Thus genuinely + new types introduced by \hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}} stay within the range + of HOL models by construction. Note that \indexref{}{command}{type\_synonym}\hyperlink{command.type-synonym}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}}}} from Isabelle/Pure merely introduces syntactic + abbreviations, without any logical significance. + + \begin{matharray}{rcl} + \indexdef{HOL}{command}{typedef}\hypertarget{command.HOL.typedef}{\hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ + \end{matharray} + + \begin{railoutput} +\rail@begin{2}{} +\rail@term{\hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}}[] \rail@bar \rail@nextbar{1} -\rail@nont{\hyperlink{syntax.parname}{\mbox{\isa{parname}}}}[] +\rail@nont{\isa{alt{\isaliteral{5F}{\isacharunderscore}}name}}[] \rail@endbar -\rail@nont{\hyperlink{syntax.typespec}{\mbox{\isa{typespec}}}}[] +\rail@nont{\isa{abs{\isaliteral{5F}{\isacharunderscore}}type}}[] +\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] +\rail@nont{\isa{rep{\isaliteral{5F}{\isacharunderscore}}set}}[] +\rail@end +\rail@begin{3}{\isa{alt{\isaliteral{5F}{\isacharunderscore}}name}} +\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] +\rail@bar +\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] +\rail@nextbar{1} +\rail@term{\isa{\isakeyword{open}}}[] +\rail@nextbar{2} +\rail@term{\isa{\isakeyword{open}}}[] +\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] +\rail@endbar +\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] +\rail@end +\rail@begin{2}{\isa{abs{\isaliteral{5F}{\isacharunderscore}}type}} +\rail@nont{\hyperlink{syntax.typespec-sorts}{\mbox{\isa{typespec{\isaliteral{5F}{\isacharunderscore}}sorts}}}}[] \rail@bar \rail@nextbar{1} \rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] \rail@endbar -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@plus -\rail@nont{\isa{cons}}[] -\rail@nextplus{1} -\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] -\rail@endplus \rail@end -\rail@begin{2}{\isa{cons}} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@endplus +\rail@begin{2}{\isa{rep{\isaliteral{5F}{\isacharunderscore}}set}} +\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] \rail@bar \rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] +\rail@term{\isa{\isakeyword{morphisms}}}[] +\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] +\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] \rail@endbar \rail@end \end{railoutput} @@ -468,25 +1545,162 @@ \begin{description} - \item \hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}} defines inductive datatypes in - HOL. + \item \hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{3D}{\isacharequal}}\ A{\isaliteral{22}{\isachardoublequote}}} + axiomatizes a type definition in the background theory of the + current context, depending on a non-emptiness result of the set + \isa{A} that needs to be proven here. The set \isa{A} may + contain type variables \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} as specified on the LHS, + but no term variables. + + Even though a local theory specification, the newly introduced type + constructor cannot depend on parameters or assumptions of the + context: this is structurally impossible in HOL. In contrast, the + non-emptiness proof may use local assumptions in unusual situations, + which could result in different interpretations in target contexts: + the meaning of the bijection between the representing set \isa{A} + and the new type \isa{t} may then change in different application + contexts. + + By default, \hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}} defines both a type + constructor \isa{t} for the new type, and a term constant \isa{t} for the representing set within the old type. Use the ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}open{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' option to suppress a separate constant definition + altogether. The injection from type to set is called \isa{Rep{\isaliteral{5F}{\isacharunderscore}}t}, + its inverse \isa{Abs{\isaliteral{5F}{\isacharunderscore}}t}, unless explicit \hyperlink{keyword.HOL.morphisms}{\mbox{\isa{\isakeyword{morphisms}}}} specification provides alternative names. - \item \hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}} represents existing types as - inductive ones, generating the standard infrastructure of derived - concepts (primitive recursion etc.). + The core axiomatization uses the locale predicate \isa{type{\isaliteral{5F}{\isacharunderscore}}definition} as defined in Isabelle/HOL. Various basic + consequences of that are instantiated accordingly, re-using the + locale facts with names derived from the new type constructor. Thus + the generic \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep} is turned into the specific + \isa{{\isaliteral{22}{\isachardoublequote}}Rep{\isaliteral{5F}{\isacharunderscore}}t{\isaliteral{22}{\isachardoublequote}}}, for example. + + Theorems \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep}, \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep{\isaliteral{5F}{\isacharunderscore}}inverse}, and \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Abs{\isaliteral{5F}{\isacharunderscore}}inverse} + provide the most basic characterization as a corresponding + injection/surjection pair (in both directions). The derived rules + \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep{\isaliteral{5F}{\isacharunderscore}}inject} and \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Abs{\isaliteral{5F}{\isacharunderscore}}inject} provide a more convenient version of + injectivity, suitable for automated proof tools (e.g.\ in + declarations involving \hyperlink{attribute.simp}{\mbox{\isa{simp}}} or \hyperlink{attribute.iff}{\mbox{\isa{iff}}}). + Furthermore, the rules \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep{\isaliteral{5F}{\isacharunderscore}}cases}~/ \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep{\isaliteral{5F}{\isacharunderscore}}induct}, and \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Abs{\isaliteral{5F}{\isacharunderscore}}cases}~/ + \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Abs{\isaliteral{5F}{\isacharunderscore}}induct} provide alternative views on + surjectivity. These rules are already declared as set or type rules + for the generic \hyperlink{method.cases}{\mbox{\isa{cases}}} and \hyperlink{method.induct}{\mbox{\isa{induct}}} methods, + respectively. + + An alternative name for the set definition (and other derived + entities) may be specified in parentheses; the default is to use + \isa{t} directly. \end{description} - The induction and exhaustion theorems generated provide case names - according to the constructors involved, while parameters are named - after the types (see also \secref{sec:cases-induct}). + \begin{warn} + If you introduce a new type axiomatically, i.e.\ via \indexref{}{command}{typedecl}\hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}} and \indexref{}{command}{axiomatization}\hyperlink{command.axiomatization}{\mbox{\isa{\isacommand{axiomatization}}}}, the minimum requirement + is that it has a non-empty model, to avoid immediate collapse of the + HOL logic. Moreover, one needs to demonstrate that the + interpretation of such free-form axiomatizations can coexist with + that of the regular \indexdef{}{command}{typedef}\hypertarget{command.typedef}{\hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}}} scheme, and any extension + that other people might have introduced elsewhere (e.g.\ in HOLCF + \cite{MuellerNvOS99}). + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Examples% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Type definitions permit the introduction of abstract data + types in a safe way, namely by providing models based on already + existing types. Given some abstract axiomatic description \isa{P} + of a type, this involves two steps: + + \begin{enumerate} + + \item Find an appropriate type \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} and subset \isa{A} which + has the desired properties \isa{P}, and make a type definition + based on this representation. + + \item Prove that \isa{P} holds for \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} by lifting \isa{P} + from the representation. + + \end{enumerate} + + You can later forget about the representation and work solely in + terms of the abstract properties \isa{P}. - See \cite{isabelle-HOL} for more details on datatypes, but beware of - the old-style theory syntax being used there! Apart from proper - proof methods for case-analysis and induction, there are also - emulations of ML tactics \hyperlink{method.HOL.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}} and \hyperlink{method.HOL.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}} available, see \secref{sec:hol-induct-tac}; these admit - to refer directly to the internal structure of subgoals (including - internally bound parameters).% + \medskip The following trivial example pulls a three-element type + into existence within the formal logical environment of HOL.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{typedef}\isamarkupfalse% +\ three\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ True{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ False{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}False{\isaliteral{2C}{\isacharcomma}}\ True{\isaliteral{29}{\isacharparenright}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{definition}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}One\ {\isaliteral{3D}{\isacharequal}}\ Abs{\isaliteral{5F}{\isacharunderscore}}three\ {\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ True{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isacommand{definition}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}Two\ {\isaliteral{3D}{\isacharequal}}\ Abs{\isaliteral{5F}{\isacharunderscore}}three\ {\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ False{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isacommand{definition}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}Three\ {\isaliteral{3D}{\isacharequal}}\ Abs{\isaliteral{5F}{\isacharunderscore}}three\ {\isaliteral{28}{\isacharparenleft}}False{\isaliteral{2C}{\isacharcomma}}\ True{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ three{\isaliteral{5F}{\isacharunderscore}}distinct{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}One\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Two{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}One\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Three{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}Two\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Three{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ One{\isaliteral{5F}{\isacharunderscore}}def\ Two{\isaliteral{5F}{\isacharunderscore}}def\ Three{\isaliteral{5F}{\isacharunderscore}}def\ Abs{\isaliteral{5F}{\isacharunderscore}}three{\isaliteral{5F}{\isacharunderscore}}inject\ three{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ three{\isaliteral{5F}{\isacharunderscore}}cases{\isaliteral{3A}{\isacharcolon}}\isanewline +\ \ \isakeyword{fixes}\ x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ three\ \isakeyword{obtains}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ One{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ Two{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ Three{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}cases\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}auto\ simp{\isaliteral{3A}{\isacharcolon}}\ One{\isaliteral{5F}{\isacharunderscore}}def\ Two{\isaliteral{5F}{\isacharunderscore}}def\ Three{\isaliteral{5F}{\isacharunderscore}}def\ Abs{\isaliteral{5F}{\isacharunderscore}}three{\isaliteral{5F}{\isacharunderscore}}inject\ three{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Note that such trivial constructions are better done with + derived specification mechanisms such as \hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ three{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3D}{\isacharequal}}\ One{\isaliteral{27}{\isacharprime}}\ {\isaliteral{7C}{\isacharbar}}\ Two{\isaliteral{27}{\isacharprime}}\ {\isaliteral{7C}{\isacharbar}}\ Three{\isaliteral{27}{\isacharprime}}% +\begin{isamarkuptext}% +This avoids re-doing basic definitions and proofs from the + primitive \hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}} above.% \end{isamarkuptext}% \isamarkuptrue% % @@ -540,641 +1754,6 @@ \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsection{Recursive functions \label{sec:recursion}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{primrec}\hypertarget{command.HOL.primrec}{\hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{fun}\hypertarget{command.HOL.fun}{\hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{function}\hypertarget{command.HOL.function}{\hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{termination}\hypertarget{command.HOL.termination}{\hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@term{\isa{\isakeyword{where}}}[] -\rail@nont{\isa{equations}}[] -\rail@end -\rail@begin{4}{} -\rail@bar -\rail@term{\hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{functionopts}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@cr{3} -\rail@term{\isa{\isakeyword{where}}}[] -\rail@nont{\isa{equations}}[] -\rail@end -\rail@begin{3}{\isa{equations}} -\rail@plus -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{\isa{functionopts}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@bar -\rail@term{\isa{sequential}}[] -\rail@nextbar{1} -\rail@term{\isa{domintros}}[] -\rail@endbar -\rail@nextplus{2} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}} defines primitive recursive - functions over datatypes, see also \cite{isabelle-HOL}. - - \item \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} defines functions by general - wellfounded recursion. A detailed description with examples can be - found in \cite{isabelle-function}. The function is specified by a - set of (possibly conditional) recursive equations with arbitrary - pattern matching. The command generates proof obligations for the - completeness and the compatibility of patterns. - - The defined function is considered partial, and the resulting - simplification rules (named \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}psimps{\isaliteral{22}{\isachardoublequote}}}) and induction rule - (named \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}pinduct{\isaliteral{22}{\isachardoublequote}}}) are guarded by a generated domain - predicate \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{5F}{\isacharunderscore}}dom{\isaliteral{22}{\isachardoublequote}}}. The \hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}} - command can then be used to establish that the function is total. - - \item \hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}} is a shorthand notation for ``\hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}sequential{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, followed by automated - proof attempts regarding pattern matching and termination. See - \cite{isabelle-function} for further details. - - \item \hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}~\isa{f} commences a - termination proof for the previously defined function \isa{f}. If - this is omitted, the command refers to the most recent function - definition. After the proof is closed, the recursive equations and - the induction principle is established. - - \end{description} - - Recursive definitions introduced by the \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} - command accommodate - reasoning by induction (cf.\ \secref{sec:cases-induct}): rule \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{2E}{\isachardot}}induct{\isaliteral{22}{\isachardoublequote}}} (where \isa{c} is the name of the function definition) - refers to a specific induction rule, with parameters named according - to the user-specified equations. Cases are numbered (starting from 1). - - For \hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}}, the induction principle coincides - with structural recursion on the datatype the recursion is carried - out. - - The equations provided by these packages may be referred later as - theorem list \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}simps{\isaliteral{22}{\isachardoublequote}}}, where \isa{f} is the (collective) - name of the functions defined. Individual equations may be named - explicitly as well. - - The \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} command accepts the following - options. - - \begin{description} - - \item \isa{sequential} enables a preprocessor which disambiguates - overlapping patterns by making them mutually disjoint. Earlier - equations take precedence over later ones. This allows to give the - specification in a format very similar to functional programming. - Note that the resulting simplification and induction rules - correspond to the transformed specification, not the one given - originally. This usually means that each equation given by the user - may result in several theorems. Also note that this automatic - transformation only works for ML-style datatype patterns. - - \item \isa{domintros} enables the automated generation of - introduction rules for the domain predicate. While mostly not - needed, they can be helpful in some proofs about partial functions. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Proof methods related to recursive definitions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{method}{pat\_completeness}\hypertarget{method.HOL.pat-completeness}{\hyperlink{method.HOL.pat-completeness}{\mbox{\isa{pat{\isaliteral{5F}{\isacharunderscore}}completeness}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{relation}\hypertarget{method.HOL.relation}{\hyperlink{method.HOL.relation}{\mbox{\isa{relation}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{lexicographic\_order}\hypertarget{method.HOL.lexicographic-order}{\hyperlink{method.HOL.lexicographic-order}{\mbox{\isa{lexicographic{\isaliteral{5F}{\isacharunderscore}}order}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{size\_change}\hypertarget{method.HOL.size-change}{\hyperlink{method.HOL.size-change}{\mbox{\isa{size{\isaliteral{5F}{\isacharunderscore}}change}}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\hyperlink{method.HOL.relation}{\mbox{\isa{relation}}}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.HOL.lexicographic-order}{\mbox{\isa{lexicographic{\isaliteral{5F}{\isacharunderscore}}order}}}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.HOL.size-change}{\mbox{\isa{size{\isaliteral{5F}{\isacharunderscore}}change}}}}[] -\rail@nont{\isa{orders}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] -\rail@endplus -\rail@end -\rail@begin{4}{\isa{orders}} -\rail@plus -\rail@nextplus{1} -\rail@bar -\rail@term{\isa{max}}[] -\rail@nextbar{2} -\rail@term{\isa{min}}[] -\rail@nextbar{3} -\rail@term{\isa{ms}}[] -\rail@endbar -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.HOL.pat-completeness}{\mbox{\isa{pat{\isaliteral{5F}{\isacharunderscore}}completeness}}} is a specialized method to - solve goals regarding the completeness of pattern matching, as - required by the \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} package (cf.\ - \cite{isabelle-function}). - - \item \hyperlink{method.HOL.relation}{\mbox{\isa{relation}}}~\isa{R} introduces a termination - proof using the relation \isa{R}. The resulting proof state will - contain goals expressing that \isa{R} is wellfounded, and that the - arguments of recursive calls decrease with respect to \isa{R}. - Usually, this method is used as the initial proof step of manual - termination proofs. - - \item \hyperlink{method.HOL.lexicographic-order}{\mbox{\isa{lexicographic{\isaliteral{5F}{\isacharunderscore}}order}}} attempts a fully - automated termination proof by searching for a lexicographic - combination of size measures on the arguments of the function. The - method accepts the same arguments as the \hyperlink{method.auto}{\mbox{\isa{auto}}} method, - which it uses internally to prove local descents. The same context - modifiers as for \hyperlink{method.auto}{\mbox{\isa{auto}}} are accepted, see - \secref{sec:clasimp}. - - In case of failure, extensive information is printed, which can help - to analyse the situation (cf.\ \cite{isabelle-function}). - - \item \hyperlink{method.HOL.size-change}{\mbox{\isa{size{\isaliteral{5F}{\isacharunderscore}}change}}} also works on termination goals, - using a variation of the size-change principle, together with a - graph decomposition technique (see \cite{krauss_phd} for details). - Three kinds of orders are used internally: \isa{max}, \isa{min}, - and \isa{ms} (multiset), which is only available when the theory - \isa{Multiset} is loaded. When no order kinds are given, they are - tried in order. The search for a termination proof uses SAT solving - internally. - - For local descent proofs, the same context modifiers as for \hyperlink{method.auto}{\mbox{\isa{auto}}} are accepted, see \secref{sec:clasimp}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Functions with explicit partiality% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{partial\_function}\hypertarget{command.HOL.partial-function}{\hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isaliteral{5F}{\isacharunderscore}}function}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{attribute}{partial\_function\_mono}\hypertarget{attribute.HOL.partial-function-mono}{\hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}mono}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{5}{} -\rail@term{\hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isaliteral{5F}{\isacharunderscore}}function}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@cr{3} -\rail@term{\isa{\isakeyword{where}}}[] -\rail@bar -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isaliteral{5F}{\isacharunderscore}}function}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}mode{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} defines - recursive functions based on fixpoints in complete partial - orders. No termination proof is required from the user or - constructed internally. Instead, the possibility of non-termination - is modelled explicitly in the result type, which contains an - explicit bottom element. - - Pattern matching and mutual recursion are currently not supported. - Thus, the specification consists of a single function described by a - single recursive equation. - - There are no fixed syntactic restrictions on the body of the - function, but the induced functional must be provably monotonic - wrt.\ the underlying order. The monotonicitity proof is performed - internally, and the definition is rejected when it fails. The proof - can be influenced by declaring hints using the - \hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}mono}}} attribute. - - The mandatory \isa{mode} argument specifies the mode of operation - of the command, which directly corresponds to a complete partial - order on the result type. By default, the following modes are - defined: - - \begin{description} - \item \isa{option} defines functions that map into the \isa{option} type. Here, the value \isa{None} is used to model a - non-terminating computation. Monotonicity requires that if \isa{None} is returned by a recursive call, then the overall result - must also be \isa{None}. This is best achieved through the use of - the monadic operator \isa{{\isaliteral{22}{\isachardoublequote}}Option{\isaliteral{2E}{\isachardot}}bind{\isaliteral{22}{\isachardoublequote}}}. - - \item \isa{tailrec} defines functions with an arbitrary result - type and uses the slightly degenerated partial order where \isa{{\isaliteral{22}{\isachardoublequote}}undefined{\isaliteral{22}{\isachardoublequote}}} is the bottom element. Now, monotonicity requires that - if \isa{undefined} is returned by a recursive call, then the - overall result must also be \isa{undefined}. In practice, this is - only satisfied when each recursive call is a tail call, whose result - is directly returned. Thus, this mode of operation allows the - definition of arbitrary tail-recursive functions. - \end{description} - - Experienced users may define new modes by instantiating the locale - \isa{{\isaliteral{22}{\isachardoublequote}}partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}definitions{\isaliteral{22}{\isachardoublequote}}} appropriately. - - \item \hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}mono}}} declares rules for - use in the internal monononicity proofs of partial function - definitions. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Old-style recursive function definitions (TFL)% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The old TFL commands \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} and \hyperlink{command.HOL.recdef-tc}{\mbox{\isa{\isacommand{recdef{\isaliteral{5F}{\isacharunderscore}}tc}}}} for defining recursive are mostly obsolete; \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} or \hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}} should be used instead. - - \begin{matharray}{rcl} - \indexdef{HOL}{command}{recdef}\hypertarget{command.HOL.recdef}{\hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{recdef\_tc}\hypertarget{command.HOL.recdef-tc}{\hyperlink{command.HOL.recdef-tc}{\mbox{\isa{\isacommand{recdef{\isaliteral{5F}{\isacharunderscore}}tc}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{5}{} -\rail@term{\hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{\isakeyword{permissive}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@cr{3} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{4} -\rail@endplus -\rail@bar -\rail@nextbar{4} -\rail@nont{\isa{hints}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@nont{\isa{recdeftc}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\isa{tc}}[] -\rail@end -\rail@begin{2}{\isa{hints}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{\isakeyword{hints}}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\isa{recdefmod}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{4}{\isa{recdefmod}} -\rail@bar -\rail@bar -\rail@term{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}}[] -\rail@nextbar{1} -\rail@term{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}}[] -\rail@nextbar{2} -\rail@term{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@nextbar{2} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@nextbar{3} -\rail@nont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\isa{tc}} -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} defines general well-founded - recursive functions (using the TFL package), see also - \cite{isabelle-HOL}. The ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}permissive{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' option tells - TFL to recover from failed proof attempts, returning unfinished - results. The \isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}, \isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}, and \isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf} hints refer to auxiliary rules to be used in the internal - automated proof process of TFL. Additional \hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}} - declarations (cf.\ \secref{sec:clasimp}) may be given to tune the - context of the Simplifier (cf.\ \secref{sec:simplifier}) and - Classical reasoner (cf.\ \secref{sec:classical}). - - \item \hyperlink{command.HOL.recdef-tc}{\mbox{\isa{\isacommand{recdef{\isaliteral{5F}{\isacharunderscore}}tc}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{28}{\isacharparenleft}}i{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} recommences the - proof for leftover termination condition number \isa{i} (default - 1) as generated by a \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} definition of - constant \isa{c}. - - Note that in most cases, \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} is able to finish - its internal proofs without manual intervention. - - \end{description} - - \medskip Hints for \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} may be also declared - globally, using the following attributes. - - \begin{matharray}{rcl} - \indexdef{HOL}{attribute}{recdef\_simp}\hypertarget{attribute.HOL.recdef-simp}{\hyperlink{attribute.HOL.recdef-simp}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{recdef\_cong}\hypertarget{attribute.HOL.recdef-cong}{\hyperlink{attribute.HOL.recdef-cong}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{recdef\_wf}\hypertarget{attribute.HOL.recdef-wf}{\hyperlink{attribute.HOL.recdef-wf}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{3}{} -\rail@bar -\rail@term{\hyperlink{attribute.HOL.recdef-simp}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{attribute.HOL.recdef-cong}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{attribute.HOL.recdef-wf}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@nextbar{2} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@end -\end{railoutput}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Inductive and coinductive definitions \label{sec:hol-inductive}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -An \textbf{inductive definition} specifies the least predicate (or - set) \isa{R} closed under given rules: applying a rule to elements - of \isa{R} yields a result within \isa{R}. For example, a - structural operational semantics is an inductive definition of an - evaluation relation. - - Dually, a \textbf{coinductive definition} specifies the greatest - predicate~/ set \isa{R} that is consistent with given rules: every - element of \isa{R} can be seen as arising by applying a rule to - elements of \isa{R}. An important example is using bisimulation - relations to formalise equivalence of processes and infinite data - structures. - - \medskip The HOL package is related to the ZF one, which is - described in a separate paper,\footnote{It appeared in CADE - \cite{paulson-CADE}; a longer version is distributed with Isabelle.} - which you should refer to in case of difficulties. The package is - simpler than that of ZF thanks to implicit type-checking in HOL. - The types of the (co)inductive predicates (or sets) determine the - domain of the fixedpoint definition, and the package does not have - to use inference rules for type-checking. - - \begin{matharray}{rcl} - \indexdef{HOL}{command}{inductive}\hypertarget{command.HOL.inductive}{\hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{inductive\_set}\hypertarget{command.HOL.inductive-set}{\hyperlink{command.HOL.inductive-set}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{coinductive}\hypertarget{command.HOL.coinductive}{\hyperlink{command.HOL.coinductive}{\mbox{\isa{\isacommand{coinductive}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{coinductive\_set}\hypertarget{command.HOL.coinductive-set}{\hyperlink{command.HOL.coinductive-set}{\mbox{\isa{\isacommand{coinductive{\isaliteral{5F}{\isacharunderscore}}set}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{attribute}{mono}\hypertarget{attribute.HOL.mono}{\hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{7}{} -\rail@bar -\rail@term{\hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.HOL.inductive-set}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{command.HOL.coinductive}{\mbox{\isa{\isacommand{coinductive}}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{command.HOL.coinductive-set}{\mbox{\isa{\isacommand{coinductive{\isaliteral{5F}{\isacharunderscore}}set}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{for}}}[] -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@endbar -\rail@cr{5} -\rail@bar -\rail@nextbar{6} -\rail@term{\isa{\isakeyword{where}}}[] -\rail@nont{\isa{clauses}}[] -\rail@endbar -\rail@bar -\rail@nextbar{6} -\rail@term{\isa{\isakeyword{monos}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\rail@begin{3}{\isa{clauses}} -\rail@plus -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@nextbar{2} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}} and \hyperlink{command.HOL.coinductive}{\mbox{\isa{\isacommand{coinductive}}}} define (co)inductive predicates from the - introduction rules given in the \hyperlink{keyword.where}{\mbox{\isa{\isakeyword{where}}}} part. The - optional \hyperlink{keyword.for}{\mbox{\isa{\isakeyword{for}}}} part contains a list of parameters of the - (co)inductive predicates that remain fixed throughout the - definition. The optional \hyperlink{keyword.monos}{\mbox{\isa{\isakeyword{monos}}}} section contains - \emph{monotonicity theorems}, which are required for each operator - applied to a recursive set in the introduction rules. There - \emph{must} be a theorem of the form \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ M\ A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ M\ B{\isaliteral{22}{\isachardoublequote}}}, - for each premise \isa{{\isaliteral{22}{\isachardoublequote}}M\ R\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ t{\isaliteral{22}{\isachardoublequote}}} in an introduction rule! - - \item \hyperlink{command.HOL.inductive-set}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}}}} and \hyperlink{command.HOL.coinductive-set}{\mbox{\isa{\isacommand{coinductive{\isaliteral{5F}{\isacharunderscore}}set}}}} are wrappers for to the previous commands, - allowing the definition of (co)inductive sets. - - \item \hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}} declares monotonicity rules. These - rule are involved in the automated monotonicity proof of \hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Derived rules% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Each (co)inductive definition \isa{R} adds definitions to the - theory and also proves some theorems: - - \begin{description} - - \item \isa{R{\isaliteral{2E}{\isachardot}}intros} is the list of introduction rules as proven - theorems, for the recursive predicates (or sets). The rules are - also available individually, using the names given them in the - theory file; - - \item \isa{R{\isaliteral{2E}{\isachardot}}cases} is the case analysis (or elimination) rule; - - \item \isa{R{\isaliteral{2E}{\isachardot}}induct} or \isa{R{\isaliteral{2E}{\isachardot}}coinduct} is the (co)induction - rule. - - \end{description} - - When several predicates \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} are - defined simultaneously, the list of introduction rules is called - \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5F}{\isacharunderscore}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}intros{\isaliteral{22}{\isachardoublequote}}}, the case analysis rules are - called \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2E}{\isachardot}}cases{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}cases{\isaliteral{22}{\isachardoublequote}}}, and the list - of mutual induction rules is called \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5F}{\isacharunderscore}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}inducts{\isaliteral{22}{\isachardoublequote}}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Monotonicity theorems% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Each theory contains a default set of theorems that are used in - monotonicity proofs. New rules can be added to this set via the - \hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}} attribute. The HOL theory \isa{Inductive} - shows how this is done. In general, the following monotonicity - theorems may be added: - - \begin{itemize} - - \item Theorems of the form \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ M\ A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ M\ B{\isaliteral{22}{\isachardoublequote}}}, for proving - monotonicity of inductive definitions whose introduction rules have - premises involving terms such as \isa{{\isaliteral{22}{\isachardoublequote}}M\ R\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ t{\isaliteral{22}{\isachardoublequote}}}. - - \item Monotonicity theorems for logical operators, which are of the - general form \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}. For example, in - the case of the operator \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6F723E}{\isasymor}}{\isaliteral{22}{\isachardoublequote}}}, the corresponding theorem is - \[ - \infer{\isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}} - \] - - \item De Morgan style equations for reasoning about the ``polarity'' - of expressions, e.g. - \[ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ P{\isaliteral{22}{\isachardoublequote}}} \qquad\qquad - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C616E643E}{\isasymand}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q{\isaliteral{22}{\isachardoublequote}}} - \] - - \item Equations for reducing complex operators to more primitive - ones whose monotonicity can easily be proved, e.g. - \[ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}} \qquad\qquad - \isa{{\isaliteral{22}{\isachardoublequote}}Ball\ A\ P\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P\ x{\isaliteral{22}{\isachardoublequote}}} - \] - - \end{itemize} - - %FIXME: Example of an inductive definition% -\end{isamarkuptext}% -\isamarkuptrue% -% \isamarkupsection{Arithmetic proof support% } \isamarkuptrue% @@ -1483,14 +2062,15 @@ \item \hyperlink{command.HOL.value}{\mbox{\isa{\isacommand{value}}}}~\isa{t} evaluates and prints a term; optionally \isa{modes} can be specified, which are - appended to the current print mode (see also \cite{isabelle-ref}). + appended to the current print mode; see \secref{sec:print-modes}. Internally, the evaluation is performed by registered evaluators, which are invoked sequentially until a result is returned. Alternatively a specific evaluator can be selected using square brackets; typical evaluators use the current set of code equations - to normalize and include \isa{simp} for fully symbolic evaluation - using the simplifier, \isa{nbe} for \emph{normalization by evaluation} - and \emph{code} for code generation in SML. + to normalize and include \isa{simp} for fully symbolic + evaluation using the simplifier, \isa{nbe} for + \emph{normalization by evaluation} and \emph{code} for code + generation in SML. \item \hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}} tests the current goal for counterexamples using a series of assignments for its diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/document/Inner_Syntax.tex --- a/doc-src/IsarRef/Thy/document/Inner_Syntax.tex Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/Thy/document/Inner_Syntax.tex Wed Jun 08 10:24:07 2011 +0200 @@ -151,8 +151,8 @@ \end{description} All of the diagnostic commands above admit a list of \isa{modes} - to be specified, which is appended to the current print mode (see - also \cite{isabelle-ref}). Thus the output behavior may be modified + to be specified, which is appended to the current print mode; see + \secref{sec:print-modes}. Thus the output behavior may be modified according particular print mode features. For example, \hyperlink{command.pr}{\mbox{\isa{\isacommand{pr}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}latex\ xsymbols{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} would print the current proof state with mathematical symbols and special characters represented in {\LaTeX} source, according to the Isabelle style diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/document/Introduction.tex --- a/doc-src/IsarRef/Thy/document/Introduction.tex Wed Jun 08 08:47:43 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,111 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Introduction}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Introduction\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Introduction% -} -\isamarkuptrue% -% -\isamarkupsection{Overview% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The \emph{Isabelle} system essentially provides a generic - infrastructure for building deductive systems (programmed in - Standard ML), with a special focus on interactive theorem proving in - higher-order logics. Many years ago, even end-users would refer to - certain ML functions (goal commands, tactics, tacticals etc.) to - pursue their everyday theorem proving tasks. - - In contrast \emph{Isar} provides an interpreted language environment - of its own, which has been specifically tailored for the needs of - theory and proof development. Compared to raw ML, the Isabelle/Isar - top-level provides a more robust and comfortable development - platform, with proper support for theory development graphs, managed - transactions with unlimited undo etc. The Isabelle/Isar version of - the \emph{Proof~General} user interface - \cite{proofgeneral,Aspinall:TACAS:2000} provides a decent front-end - for interactive theory and proof development in this advanced - theorem proving environment, even though it is somewhat biased - towards old-style proof scripts. - - \medskip Apart from the technical advances over bare-bones ML - programming, the main purpose of the Isar language is to provide a - conceptually different view on machine-checked proofs - \cite{Wenzel:1999:TPHOL,Wenzel-PhD}. \emph{Isar} stands for - \emph{Intelligible semi-automated reasoning}. Drawing from both the - traditions of informal mathematical proof texts and high-level - programming languages, Isar offers a versatile environment for - structured formal proof documents. Thus properly written Isar - proofs become accessible to a broader audience than unstructured - tactic scripts (which typically only provide operational information - for the machine). Writing human-readable proof texts certainly - requires some additional efforts by the writer to achieve a good - presentation, both of formal and informal parts of the text. On the - other hand, human-readable formal texts gain some value in their own - right, independently of the mechanic proof-checking process. - - Despite its grand design of structured proof texts, Isar is able to - assimilate the old tactical style as an ``improper'' sub-language. - This provides an easy upgrade path for existing tactic scripts, as - well as some means for interactive experimentation and debugging of - structured proofs. Isabelle/Isar supports a broad range of proof - styles, both readable and unreadable ones. - - \medskip The generic Isabelle/Isar framework (see - \chref{ch:isar-framework}) works reasonably well for any Isabelle - object-logic that conforms to the natural deduction view of the - Isabelle/Pure framework. Specific language elements introduced by - the major object-logics are described in \chref{ch:hol} - (Isabelle/HOL), \chref{ch:holcf} (Isabelle/HOLCF), and \chref{ch:zf} - (Isabelle/ZF). The main language elements are already provided by - the Isabelle/Pure framework. Nevertheless, examples given in the - generic parts will usually refer to Isabelle/HOL as well. - - \medskip Isar commands may be either \emph{proper} document - constructors, or \emph{improper commands}. Some proof methods and - attributes introduced later are classified as improper as well. - Improper Isar language elements, which are marked by ``\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}'' in the subsequent chapters; they are often helpful - when developing proof documents, but their use is discouraged for - the final human-readable outcome. Typical examples are diagnostic - commands that print terms or theorems according to the current - context; other commands emulate old-style tactical theorem proving.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/document/Preface.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Thy/document/Preface.tex Wed Jun 08 10:24:07 2011 +0200 @@ -0,0 +1,107 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Preface}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Preface\isanewline +\isakeyword{imports}\ Base\ Main\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupchapter{Preface% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The \emph{Isabelle} system essentially provides a generic + infrastructure for building deductive systems (programmed in + Standard ML), with a special focus on interactive theorem proving in + higher-order logics. Many years ago, even end-users would refer to + certain ML functions (goal commands, tactics, tacticals etc.) to + pursue their everyday theorem proving tasks. + + In contrast \emph{Isar} provides an interpreted language environment + of its own, which has been specifically tailored for the needs of + theory and proof development. Compared to raw ML, the Isabelle/Isar + top-level provides a more robust and comfortable development + platform, with proper support for theory development graphs, managed + transactions with unlimited undo etc. The Isabelle/Isar version of + the \emph{Proof~General} user interface + \cite{proofgeneral,Aspinall:TACAS:2000} provides a decent front-end + for interactive theory and proof development in this advanced + theorem proving environment, even though it is somewhat biased + towards old-style proof scripts. + + \medskip Apart from the technical advances over bare-bones ML + programming, the main purpose of the Isar language is to provide a + conceptually different view on machine-checked proofs + \cite{Wenzel:1999:TPHOL,Wenzel-PhD}. \emph{Isar} stands for + \emph{Intelligible semi-automated reasoning}. Drawing from both the + traditions of informal mathematical proof texts and high-level + programming languages, Isar offers a versatile environment for + structured formal proof documents. Thus properly written Isar + proofs become accessible to a broader audience than unstructured + tactic scripts (which typically only provide operational information + for the machine). Writing human-readable proof texts certainly + requires some additional efforts by the writer to achieve a good + presentation, both of formal and informal parts of the text. On the + other hand, human-readable formal texts gain some value in their own + right, independently of the mechanic proof-checking process. + + Despite its grand design of structured proof texts, Isar is able to + assimilate the old tactical style as an ``improper'' sub-language. + This provides an easy upgrade path for existing tactic scripts, as + well as some means for interactive experimentation and debugging of + structured proofs. Isabelle/Isar supports a broad range of proof + styles, both readable and unreadable ones. + + \medskip The generic Isabelle/Isar framework (see + \chref{ch:isar-framework}) works reasonably well for any Isabelle + object-logic that conforms to the natural deduction view of the + Isabelle/Pure framework. Specific language elements introduced by + the major object-logics are described in \chref{ch:hol} + (Isabelle/HOL), \chref{ch:holcf} (Isabelle/HOLCF), and \chref{ch:zf} + (Isabelle/ZF). The main language elements are already provided by + the Isabelle/Pure framework. Nevertheless, examples given in the + generic parts will usually refer to Isabelle/HOL as well. + + \medskip Isar commands may be either \emph{proper} document + constructors, or \emph{improper commands}. Some proof methods and + attributes introduced later are classified as improper as well. + Improper Isar language elements, which are marked by ``\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}'' in the subsequent chapters; they are often helpful + when developing proof documents, but their use is discouraged for + the final human-readable outcome. Typical examples are diagnostic + commands that print terms or theorems according to the current + context; other commands emulate old-style tactical theorem proving.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/document/Spec.tex --- a/doc-src/IsarRef/Thy/document/Spec.tex Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/Thy/document/Spec.tex Wed Jun 08 10:24:07 2011 +0200 @@ -18,7 +18,7 @@ % \endisadelimtheory % -\isamarkupchapter{Theory specifications% +\isamarkupchapter{Specifications% } \isamarkuptrue% % @@ -1357,7 +1357,7 @@ \ \ \ \ \ \ Thm{\isaliteral{2E}{\isachardot}}rule{\isaliteral{5F}{\isacharunderscore}}attribute\isanewline \ \ \ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}fn\ context{\isaliteral{3A}{\isacharcolon}}\ Context{\isaliteral{2E}{\isachardot}}generic\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ th{\isaliteral{3A}{\isacharcolon}}\ thm\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline \ \ \ \ \ \ \ \ \ \ let\ val\ th{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3D}{\isacharequal}}\ th\ OF\ ths\isanewline -\ \ \ \ \ \ \ \ \ \ in\ th{\isaliteral{27}{\isacharprime}}\ end{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2A7D}{\isacharverbatimclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}my\ rule{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \ \ \ \ \ \ in\ th{\isaliteral{27}{\isacharprime}}\ end{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2A7D}{\isacharverbatimclose}}\isanewline \isanewline \ \ \isacommand{attribute{\isaliteral{5F}{\isacharunderscore}}setup}\isamarkupfalse% \ my{\isaliteral{5F}{\isacharunderscore}}declaration\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline @@ -1365,7 +1365,7 @@ \ \ \ \ \ \ Thm{\isaliteral{2E}{\isachardot}}declaration{\isaliteral{5F}{\isacharunderscore}}attribute\isanewline \ \ \ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}fn\ th{\isaliteral{3A}{\isacharcolon}}\ thm\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ context{\isaliteral{3A}{\isacharcolon}}\ Context{\isaliteral{2E}{\isachardot}}generic\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline \ \ \ \ \ \ \ \ \ \ let\ val\ context{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3D}{\isacharequal}}\ context\isanewline -\ \ \ \ \ \ \ \ \ \ in\ context{\isaliteral{27}{\isacharprime}}\ end{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2A7D}{\isacharverbatimclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}my\ declaration{\isaliteral{22}{\isachardoublequoteclose}}% +\ \ \ \ \ \ \ \ \ \ in\ context{\isaliteral{27}{\isacharprime}}\ end{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2A7D}{\isacharverbatimclose}}% \endisatagML {\isafoldML}% % diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/Thy/document/Synopsis.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Thy/document/Synopsis.tex Wed Jun 08 10:24:07 2011 +0200 @@ -0,0 +1,2717 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Synopsis}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Synopsis\isanewline +\isakeyword{imports}\ Base\ Main\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupchapter{Synopsis% +} +\isamarkuptrue% +% +\isamarkupsection{Notepad% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +An Isar proof body serves as mathematical notepad to compose logical + content, consisting of types, terms, facts.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Types and terms% +} +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +Locally fixed entities:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{fix}\isamarkupfalse% +\ x\ \ \ % +\isamarkupcmt{local constant, without any type information yet% +} +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ \ % +\isamarkupcmt{variant with explicit type-constraint for subsequent use% +} +\isanewline +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ a\ b\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\ \ % +\isamarkupcmt{type assignment at first occurrence in concrete term% +} +% +\begin{isamarkuptxt}% +Definitions (non-polymorphic):% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{def}\isamarkupfalse% +\ x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{22}{\isachardoublequoteopen}}t{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{22}{\isachardoublequoteclose}}% +\begin{isamarkuptxt}% +Abbreviations (polymorphic):% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{let}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}f\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{22}{\isachardoublequoteclose}}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\ \ \isacommand{term}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{3F}{\isacharquery}}f\ {\isaliteral{3F}{\isacharquery}}f{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +Notation:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{write}\isamarkupfalse% +\ x\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsection{Facts% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +A fact is a simultaneous list of theorems.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Producing facts% +} +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +Via assumption (``lambda''):% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{assume}\isamarkupfalse% +\ a{\isaliteral{3A}{\isacharcolon}}\ A% +\begin{isamarkuptxt}% +Via proof (``let''):% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{have}\isamarkupfalse% +\ b{\isaliteral{3A}{\isacharcolon}}\ B\ \isacommand{sorry}\isamarkupfalse% +% +\begin{isamarkuptxt}% +Via abbreviation (``let''):% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{note}\isamarkupfalse% +\ c\ {\isaliteral{3D}{\isacharequal}}\ a\ b% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsubsection{Referencing facts% +} +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +Via explicit name:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{assume}\isamarkupfalse% +\ a{\isaliteral{3A}{\isacharcolon}}\ A\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ a% +\begin{isamarkuptxt}% +Via implicit name:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{assume}\isamarkupfalse% +\ A\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ this% +\begin{isamarkuptxt}% +Via literal proposition (unification with results from the proof text):% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{assume}\isamarkupfalse% +\ A\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}A{\isaliteral{60}{\isacharbackquoteclose}}\isanewline +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}B\ a{\isaliteral{60}{\isacharbackquoteclose}}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}B\ b{\isaliteral{60}{\isacharbackquoteclose}}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsubsection{Manipulating facts% +} +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +Instantiation:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{assume}\isamarkupfalse% +\ a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ a\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ a\ {\isaliteral{5B}{\isacharbrackleft}}of\ b{\isaliteral{5D}{\isacharbrackright}}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ a\ {\isaliteral{5B}{\isacharbrackleft}}\isakeyword{where}\ x\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{5D}{\isacharbrackright}}% +\begin{isamarkuptxt}% +Backchaining:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isadigit{1}}{\isaliteral{3A}{\isacharcolon}}\ A\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isadigit{2}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isadigit{2}}\ {\isaliteral{5B}{\isacharbrackleft}}OF\ {\isadigit{1}}{\isaliteral{5D}{\isacharbrackright}}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isadigit{1}}\ {\isaliteral{5B}{\isacharbrackleft}}THEN\ {\isadigit{2}}{\isaliteral{5D}{\isacharbrackright}}% +\begin{isamarkuptxt}% +Symmetric results:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ this\ {\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}\isanewline +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ this\ {\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}% +\begin{isamarkuptxt}% +Adhoc-simplification (take care!):% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{40}{\isacharat}}\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ this\ {\isaliteral{5B}{\isacharbrackleft}}simplified{\isaliteral{5D}{\isacharbrackright}}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsubsection{Projections% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isar facts consist of multiple theorems. There is notation to project + interval ranges.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{assume}\isamarkupfalse% +\ stuff{\isaliteral{3A}{\isacharcolon}}\ A\ B\ C\ D\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ stuff{\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ stuff{\isaliteral{28}{\isacharparenleft}}{\isadigit{2}}{\isaliteral{2D}{\isacharminus}}{\isadigit{3}}{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ stuff{\isaliteral{28}{\isacharparenleft}}{\isadigit{2}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{29}{\isacharparenright}}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsubsection{Naming conventions% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{itemize} + + \item Lower-case identifiers are usually preferred. + + \item Facts can be named after the main term within the proposition. + + \item Facts should \emph{not} be named after the command that + introduced them (\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}, \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}). This is + misleading and hard to maintain. + + \item Natural numbers can be used as ``meaningless'' names (more + appropriate than \isa{{\isaliteral{22}{\isachardoublequote}}a{\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}a{\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} etc.) + + \item Symbolic identifiers are supported (e.g. \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}). + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Block structure% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The formal notepad is block structured. The fact produced by the last + entry of a block is exported into the outer context.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ a{\isaliteral{3A}{\isacharcolon}}\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ b{\isaliteral{3A}{\isacharcolon}}\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{note}\isamarkupfalse% +\ a\ b\isanewline +\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ this\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}A{\isaliteral{60}{\isacharbackquoteclose}}\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}B{\isaliteral{60}{\isacharbackquoteclose}}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\begin{isamarkuptext}% +Explicit blocks as well as implicit blocks of nested goal + statements (e.g.\ \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}) automatically introduce one extra + pair of parentheses in reserve. The \hyperlink{command.next}{\mbox{\isa{\isacommand{next}}}} command allows + to ``jump'' between these sub-blocks.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\isanewline +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ a{\isaliteral{3A}{\isacharcolon}}\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ b{\isaliteral{3A}{\isacharcolon}}\ B\isanewline +\ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{2D}{\isacharminus}}\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ c{\isaliteral{3A}{\isacharcolon}}\ C\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ d{\isaliteral{3A}{\isacharcolon}}\ D\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +% +\begin{isamarkuptxt}% +Alternative version with explicit parentheses everywhere:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ a{\isaliteral{3A}{\isacharcolon}}\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ b{\isaliteral{3A}{\isacharcolon}}\ B\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{2D}{\isacharminus}}\isanewline +\ \ \ \ \ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ c{\isaliteral{3A}{\isacharcolon}}\ C\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ d{\isaliteral{3A}{\isacharcolon}}\ D\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{end}\isamarkupfalse% +% +\isamarkupsection{Calculational reasoning \label{sec:calculations-synopsis}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +For example, see \verb|~~/src/HOL/Isar_Examples/Group.thy|.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Special names in Isar proofs% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{itemize} + + \item term \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}thesis{\isaliteral{22}{\isachardoublequote}}} --- the main conclusion of the + innermost pending claim + + \item term \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} --- the argument of the last explicitly + stated result (for infix application this is the right-hand side) + + \item fact \isa{{\isaliteral{22}{\isachardoublequote}}this{\isaliteral{22}{\isachardoublequote}}} --- the last result produced in the text + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{2D}{\isacharminus}}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\ \ \ \ \isacommand{term}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}thesis\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{show}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\ \ \ \ \isacommand{term}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}thesis\ \ % +\isamarkupcmt{static!% +} +\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\ \ \isacommand{term}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{thm}\isamarkupfalse% +\ this\isanewline +\isacommand{end}\isamarkupfalse% +% +\begin{isamarkuptext}% +Calculational reasoning maintains the special fact called + ``\isa{calculation}'' in the background. Certain language + elements combine primary \isa{this} with secondary \isa{calculation}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Transitive chains% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The Idea is to combine \isa{this} and \isa{calculation} + via typical \isa{trans} rules (see also \hyperlink{command.print-trans-rules}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}trans{\isaliteral{5F}{\isacharunderscore}}rules}}}}):% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{thm}\isamarkupfalse% +\ trans\isanewline +\isacommand{thm}\isamarkupfalse% +\ less{\isaliteral{5F}{\isacharunderscore}}trans\isanewline +\isacommand{thm}\isamarkupfalse% +\ less{\isaliteral{5F}{\isacharunderscore}}le{\isaliteral{5F}{\isacharunderscore}}trans\isanewline +\isanewline +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +Plain bottom-up calculation:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}b\ {\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}c\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{finally}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +% +\begin{isamarkuptxt}% +Variant using the \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} abbreviation:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{finally}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +% +\begin{isamarkuptxt}% +Top-down version with explicit claim at the head:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{2D}{\isacharminus}}\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{also}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{also}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{finally}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isacommand{next}\isamarkupfalse% +% +\begin{isamarkuptxt}% +Mixed inequalities (require suitable base type):% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{fix}\isamarkupfalse% +\ a\ b\ c\ d\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3C}{\isacharless}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}b{\isaliteral{5C3C6C653E}{\isasymle}}\ c{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}c\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{finally}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3C}{\isacharless}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsubsection{Notes% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{itemize} + + \item The notion of \isa{trans} rule is very general due to the + flexibility of Isabelle/Pure rule composition. + + \item User applications may declare there own rules, with some care + about the operational details of higher-order unification. + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Degenerate calculations and bigstep reasoning% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The Idea is to append \isa{this} to \isa{calculation}, + without rule composition.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\begin{isamarkuptxt}% +A vacuous proof:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{have}\isamarkupfalse% +\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{moreover}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{moreover}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ C\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{ultimately}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ A\ \isakeyword{and}\ B\ \isakeyword{and}\ C\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isacommand{next}\isamarkupfalse% +% +\begin{isamarkuptxt}% +Slightly more content (trivial bigstep reasoning):% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{have}\isamarkupfalse% +\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{moreover}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{moreover}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ C\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{ultimately}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C616E643E}{\isasymand}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% +\ blast\isanewline +\isacommand{next}\isamarkupfalse% +% +\begin{isamarkuptxt}% +More ambitious bigstep reasoning involving structured results:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B\ {\isaliteral{5C3C6F723E}{\isasymor}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{moreover}\isamarkupfalse% +\isanewline +\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\ \isacommand{assume}\isamarkupfalse% +\ A\ \isacommand{have}\isamarkupfalse% +\ R\ \isacommand{sorry}\isamarkupfalse% +\ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \isacommand{moreover}\isamarkupfalse% +\isanewline +\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\ \isacommand{assume}\isamarkupfalse% +\ B\ \isacommand{have}\isamarkupfalse% +\ R\ \isacommand{sorry}\isamarkupfalse% +\ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \isacommand{moreover}\isamarkupfalse% +\isanewline +\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\ \isacommand{assume}\isamarkupfalse% +\ C\ \isacommand{have}\isamarkupfalse% +\ R\ \isacommand{sorry}\isamarkupfalse% +\ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \isacommand{ultimately}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ R\ \isacommand{by}\isamarkupfalse% +\ blast\ \ % +\isamarkupcmt{``big-bang integration'' of proof blocks (occasionally fragile)% +} +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{end}\isamarkupfalse% +% +\isamarkupsection{Induction% +} +\isamarkuptrue% +% +\isamarkupsubsection{Induction as Natural Deduction% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In principle, induction is just a special case of Natural + Deduction (see also \secref{sec:natural-deduction-synopsis}). For + example:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{thm}\isamarkupfalse% +\ nat{\isaliteral{2E}{\isachardot}}induct\isanewline +\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% +\ nat{\isaliteral{2E}{\isachardot}}induct\isanewline +\isanewline +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}rule\ nat{\isaliteral{2E}{\isachardot}}induct{\isaliteral{29}{\isacharparenright}}\ \ % +\isamarkupcmt{fragile rule application!% +} +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\begin{isamarkuptext}% +In practice, much more proof infrastructure is required. + + The proof method \hyperlink{method.induct}{\mbox{\isa{induct}}} provides: + \begin{itemize} + + \item implicit rule selection and robust instantiation + + \item context elements via symbolic case names + + \item support for rule-structured induction statements, with local + parameters, premises, etc. + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ n{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isadigit{0}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}case\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ Suc{\isaliteral{2E}{\isachardot}}hyps\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}case\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsubsection{Example% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The subsequent example combines the following proof patterns: + \begin{itemize} + + \item outermost induction (over the datatype structure of natural + numbers), to decompose the proof problem in top-down manner + + \item calculational reasoning (\secref{sec:calculations-synopsis}) + to compose the result in each case + + \item solving local claims within the calculation by simplification + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\isanewline +\ \ \isakeyword{fixes}\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline +\ \ \isakeyword{shows}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C53756D3E}{\isasymSum}}i{\isaliteral{3D}{\isacharequal}}{\isadigit{0}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}n{\isaliteral{2E}{\isachardot}}\ i{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ n\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ div\ {\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ n{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \isacommand{case}\isamarkupfalse% +\ {\isadigit{0}}\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C53756D3E}{\isasymSum}}i{\isaliteral{3D}{\isacharequal}}{\isadigit{0}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isadigit{0}}{\isaliteral{2E}{\isachardot}}\ i{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{0}}\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ div\ {\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}case\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{case}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C53756D3E}{\isasymSum}}i{\isaliteral{3D}{\isacharequal}}{\isadigit{0}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}Suc\ n{\isaliteral{2E}{\isachardot}}\ i{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C53756D3E}{\isasymSum}}i{\isaliteral{3D}{\isacharequal}}{\isadigit{0}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}n{\isaliteral{2E}{\isachardot}}\ i{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ n\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ div\ {\isadigit{2}}\ {\isaliteral{2B}{\isacharplus}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ Suc{\isaliteral{2E}{\isachardot}}hyps{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ div\ {\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{also}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}Suc\ n\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{28}{\isacharparenleft}}Suc\ n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ div\ {\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \isacommand{finally}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}case\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +This demonstrates how induction proofs can be done without + having to consider the raw Natural Deduction structure.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Induction with local parameters and premises% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Idea: Pure rule statements are passed through the induction + rule. This achieves convenient proof patterns, thanks to some + internal trickery in the \hyperlink{method.induct}{\mbox{\isa{induct}}} method. + + Important: Using compact HOL formulae with \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}{\isaliteral{2F}{\isacharslash}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} is a + well-known anti-pattern! It would produce useless formal noise.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ P\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ Q\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ n{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isadigit{0}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}P\ n{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ n{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isadigit{0}}\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isadigit{0}}{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}A\ n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ n{\isaliteral{60}{\isacharbackquoteclose}}\isanewline +\ \ \ \ \ \ \isakeyword{and}\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ Q\ x\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ n{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isadigit{0}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}Q\ x\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ Q\ x\ n{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}Q\ x\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +% +\begin{isamarkuptxt}% +Local quantification admits arbitrary instances:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \ \ \isacommand{note}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}Q\ a\ n{\isaliteral{60}{\isacharbackquoteclose}}\ \isakeyword{and}\ {\isaliteral{60}{\isacharbackquoteopen}}Q\ b\ n{\isaliteral{60}{\isacharbackquoteclose}}\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsection{Implicit induction context% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The \hyperlink{method.induct}{\mbox{\isa{induct}}} method can isolate local parameters and + premises directly from the given statement. This is convenient in + practical applications, but requires some understanding of what is + going on internally (as explained above).% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ Q\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ x\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}Q\ x\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ n\ arbitrary{\isaliteral{3A}{\isacharcolon}}\ x{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isadigit{0}}\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}A\ x\ {\isadigit{0}}{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}Q\ x\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ A\ x\ n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q\ x\ n{\isaliteral{60}{\isacharbackquoteclose}}\ \ % +\isamarkupcmt{arbitrary instances can be produced here% +} +\isanewline +\ \ \ \ \ \ \isakeyword{and}\ {\isaliteral{60}{\isacharbackquoteopen}}A\ x\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}Q\ x\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsection{Advanced induction with term definitions% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Induction over subexpressions of a certain shape are delicate + to formalize. The Isar \hyperlink{method.induct}{\mbox{\isa{induct}}} method provides + infrastructure for this. + + Idea: sub-expressions of the problem are turned into a defined + induction variable; often accompanied with fixing of auxiliary + parameters in the original expression.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ a\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{fix}\isamarkupfalse% +\ A\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}induct\ {\isaliteral{22}{\isachardoublequoteopen}}a\ x{\isaliteral{22}{\isachardoublequoteclose}}\ arbitrary{\isaliteral{3A}{\isacharcolon}}\ x{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isadigit{0}}\isanewline +\ \ \ \ \isacommand{note}\isamarkupfalse% +\ prem\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{60}{\isacharbackquoteclose}}\isanewline +\ \ \ \ \ \ \isakeyword{and}\ defn\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{60}{\isacharbackquoteopen}}{\isadigit{0}}\ {\isaliteral{3D}{\isacharequal}}\ a\ x{\isaliteral{60}{\isacharbackquoteclose}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{note}\isamarkupfalse% +\ hyp\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ n\ {\isaliteral{3D}{\isacharequal}}\ a\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{60}{\isacharbackquoteclose}}\isanewline +\ \ \ \ \ \ \isakeyword{and}\ prem\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{60}{\isacharbackquoteclose}}\isanewline +\ \ \ \ \ \ \isakeyword{and}\ defn\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{60}{\isacharbackquoteopen}}Suc\ n\ {\isaliteral{3D}{\isacharequal}}\ a\ x{\isaliteral{60}{\isacharbackquoteclose}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\isamarkupsection{Natural Deduction \label{sec:natural-deduction-synopsis}% +} +\isamarkuptrue% +% +\isamarkupsubsection{Rule statements% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle/Pure ``theorems'' are always natural deduction rules, + which sometimes happen to consist of a conclusion only. + + The framework connectives \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} indicate the + rule structure declaratively. For example:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{thm}\isamarkupfalse% +\ conjI\isanewline +\isacommand{thm}\isamarkupfalse% +\ impI\isanewline +\isacommand{thm}\isamarkupfalse% +\ nat{\isaliteral{2E}{\isachardot}}induct% +\begin{isamarkuptext}% +The object-logic is embedded into the Pure framework via an implicit + derivability judgment \isa{{\isaliteral{22}{\isachardoublequote}}Trueprop\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ bool\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ prop{\isaliteral{22}{\isachardoublequote}}}. + + Thus any HOL formulae appears atomic to the Pure framework, while + the rule structure outlines the corresponding proof pattern. + + This can be made explicit as follows:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{write}\isamarkupfalse% +\ Trueprop\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}Tr{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\ \ \isacommand{thm}\isamarkupfalse% +\ conjI\isanewline +\ \ \isacommand{thm}\isamarkupfalse% +\ impI\isanewline +\ \ \isacommand{thm}\isamarkupfalse% +\ nat{\isaliteral{2E}{\isachardot}}induct\isanewline +\isacommand{end}\isamarkupfalse% +% +\begin{isamarkuptext}% +Isar provides first-class notation for rule statements as follows.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% +\ conjI\isanewline +\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% +\ impI\isanewline +\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% +\ nat{\isaliteral{2E}{\isachardot}}induct% +\isamarkupsubsubsection{Examples% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Introductions and eliminations of some standard connectives of + the object-logic can be written as rule statements as follows. (The + proof ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\hyperlink{method.blast}{\mbox{\isa{blast}}}'' serves as sanity check.)% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ False{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C616E643E}{\isasymand}}\ Q{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{5C3C616E643E}{\isasymand}}\ Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ x{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C696E7465723E}{\isasyminter}}\ B{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C696E7465723E}{\isasyminter}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ B{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ B{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isacommand{lemma}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequoteclose}}% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Isar context elements% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We derive some results out of the blue, using Isar context + elements and some explicit blocks. This illustrates their meaning + wrt.\ Pure connectives, without goal states getting in the way.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ A\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{def}\isamarkupfalse% +\ x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}B\ t{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% +\ fact\isanewline +\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{obtain}\isamarkupfalse% +\ x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{have}\isamarkupfalse% +\ C\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ C\ \isacommand{by}\isamarkupfalse% +\ fact% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsection{Pure rule composition% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The Pure framework provides means for: + + \begin{itemize} + + \item backward-chaining of rules by \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} + + \item closing of branches by \hyperlink{inference.assumption}{\mbox{\isa{assumption}}} + + \end{itemize} + + Both principles involve higher-order unification of \isa{{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}}-terms + modulo \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{5C3C626574613E}{\isasymbeta}}{\isaliteral{5C3C6574613E}{\isasymeta}}{\isaliteral{22}{\isachardoublequote}}}-equivalence (cf.\ Huet and Miller).% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{assume}\isamarkupfalse% +\ a{\isaliteral{3A}{\isacharcolon}}\ A\ \isakeyword{and}\ b{\isaliteral{3A}{\isacharcolon}}\ B% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\ \ \isacommand{thm}\isamarkupfalse% +\ conjI\isanewline +\ \ \isacommand{thm}\isamarkupfalse% +\ conjI\ {\isaliteral{5B}{\isacharbrackleft}}of\ A\ B{\isaliteral{5D}{\isacharbrackright}}\ \ % +\isamarkupcmt{instantiation% +} +\isanewline +\ \ \isacommand{thm}\isamarkupfalse% +\ conjI\ {\isaliteral{5B}{\isacharbrackleft}}of\ A\ B{\isaliteral{2C}{\isacharcomma}}\ OF\ a\ b{\isaliteral{5D}{\isacharbrackright}}\ \ % +\isamarkupcmt{instantiation and composition% +} +\isanewline +\ \ \isacommand{thm}\isamarkupfalse% +\ conjI\ {\isaliteral{5B}{\isacharbrackleft}}OF\ a\ b{\isaliteral{5D}{\isacharbrackright}}\ \ % +\isamarkupcmt{composition via unification (trivial)% +} +\isanewline +\ \ \isacommand{thm}\isamarkupfalse% +\ conjI\ {\isaliteral{5B}{\isacharbrackleft}}OF\ {\isaliteral{60}{\isacharbackquoteopen}}A{\isaliteral{60}{\isacharbackquoteclose}}\ {\isaliteral{60}{\isacharbackquoteopen}}B{\isaliteral{60}{\isacharbackquoteclose}}{\isaliteral{5D}{\isacharbrackright}}\isanewline +\isanewline +\ \ \isacommand{thm}\isamarkupfalse% +\ conjI\ {\isaliteral{5B}{\isacharbrackleft}}OF\ disjI{\isadigit{1}}{\isaliteral{5D}{\isacharbrackright}}\isanewline +\isacommand{end}\isamarkupfalse% +% +\begin{isamarkuptext}% +Note: Low-level rule composition is tedious and leads to + unreadable~/ unmaintainable expressions in the text.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Structured backward reasoning% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Idea: Canonical proof decomposition via \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~/ + \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~/ \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}, where the body produces a + natural deduction rule to refine some goal.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ A\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ A\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{2D}{\isacharminus}}\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ A\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{2D}{\isacharminus}}\isanewline +\ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\isanewline +\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\ % +\isamarkupcmt{implicit block structure made explicit% +} +\isanewline +\ \ \ \ \isacommand{note}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ A\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ x{\isaliteral{60}{\isacharbackquoteclose}}\isanewline +\ \ \ \ \ \ % +\isamarkupcmt{side exit for the resulting rule% +} +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsection{Structured rule application% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Idea: Previous facts and new claims are composed with a rule from + the context (or background library).% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{assume}\isamarkupfalse% +\ r{\isadigit{1}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\ \ % +\isamarkupcmt{simple rule (Horn clause)% +} +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ A\ \isacommand{sorry}\isamarkupfalse% +\ \ % +\isamarkupcmt{prefix of facts via outer sub-proof% +} +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ C\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}rule\ r{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ B\ \isacommand{sorry}\isamarkupfalse% +\ \ % +\isamarkupcmt{remaining rule premises via inner sub-proof% +} +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ C\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}rule\ r{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ A\ \isakeyword{and}\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ C\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}rule\ r{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ A\ \isakeyword{and}\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ C\ \isacommand{by}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}rule\ r{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\isanewline +\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{assume}\isamarkupfalse% +\ r{\isadigit{2}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B{\isadigit{1}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isadigit{2}}\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\ \ % +\isamarkupcmt{nested rule% +} +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ C\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}rule\ r{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}B{\isadigit{1}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}B{\isadigit{2}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptxt}% +The compound rule premise \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B{\isadigit{1}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isadigit{2}}\ x{\isaliteral{22}{\isachardoublequote}}} is better + addressed via \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~/ \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~/ \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} + in the nested proof body.% +\end{isamarkuptxt}% +\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsection{Example: predicate logic% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Using the above principles, standard introduction and elimination proofs + of predicate logic connectives of HOL work as follows.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ B\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ C\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ A\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ C\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ B\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ C\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ A\ \isakeyword{and}\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ A\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ B\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ False\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ A\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ True\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ A\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ False\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ B\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ C\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ a\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}P\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ C\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptxt}% +Less awkward version using \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}:% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ a\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}P\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\begin{isamarkuptext}% +Further variations to illustrate Isar sub-proofs involving + \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ \ % +\isamarkupcmt{two strictly isolated subproofs% +} +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ \ % +\isamarkupcmt{one simultaneous sub-proof% +} +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ A\ \isakeyword{and}\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ \ % +\isamarkupcmt{two subproofs in the same context% +} +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ \ % +\isamarkupcmt{swapped order% +} +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ B\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ \ % +\isamarkupcmt{sequential subproofs% +} +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ A\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ B\ \isacommand{using}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}A{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsubsection{Example: set-theoretic operators% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +There is nothing special about logical connectives (\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616E643E}{\isasymand}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6F723E}{\isasymor}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}{\isaliteral{22}{\isachardoublequote}}} etc.). Operators from + set-theory or lattice-theory for analogously. It is only a matter + of rule declarations in the library; rules can be also specified + explicitly.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C696E7465723E}{\isasyminter}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ C\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ C\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ C\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\isacommand{next}\isamarkupfalse% +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C496E7465723E}{\isasymInter}}A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ a\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C496E7465723E}{\isasymInter}}A{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C556E696F6E3E}{\isasymUnion}}A{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C556E696F6E3E}{\isasymUnion}}A{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{obtain}\isamarkupfalse% +\ a\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\isamarkupsection{Generalized elimination and cases% +} +\isamarkuptrue% +% +\isamarkupsubsection{General elimination rules% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The general format of elimination rules is illustrated by the + following typical representatives:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{thm}\isamarkupfalse% +\ exE\ \ \ \ \ % +\isamarkupcmt{local parameter% +} +\isanewline +\isacommand{thm}\isamarkupfalse% +\ conjE\ \ \ % +\isamarkupcmt{local premises% +} +\isanewline +\isacommand{thm}\isamarkupfalse% +\ disjE\ \ \ % +\isamarkupcmt{split into cases% +} +% +\begin{isamarkuptext}% +Combining these characteristics leads to the following general scheme + for elimination rules with cases: + + \begin{itemize} + + \item prefix of assumptions (or ``major premises'') + + \item one or more cases that enable to establish the main conclusion + in an augmented context + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{assume}\isamarkupfalse% +\ r{\isaliteral{3A}{\isacharcolon}}\isanewline +\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}A{\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A{\isadigit{2}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}\ assumptions\ {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ B{\isadigit{1}}\ x\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isadigit{1}}\ x\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}\ case\ {\isadigit{1}}\ {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ B{\isadigit{2}}\ x\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isadigit{2}}\ x\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}\ case\ {\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \ \ R\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}\ main\ conclusion\ {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ A{\isadigit{1}}\ \isakeyword{and}\ A{\isadigit{2}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ R\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}rule\ r{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\ y\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}B{\isadigit{1}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}C{\isadigit{1}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\ y\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isaliteral{22}{\isachardoublequoteopen}}B{\isadigit{2}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}C{\isadigit{2}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\begin{isamarkuptext}% +Here \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}thesis{\isaliteral{22}{\isachardoublequote}}} is used to refer to the unchanged goal + statement.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Rules with cases% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Applying an elimination rule to some goal, leaves that unchanged + but allows to augment the context in the sub-proof of each case. + + Isar provides some infrastructure to support this: + + \begin{itemize} + + \item native language elements to state eliminations + + \item symbolic case names + + \item method \hyperlink{method.cases}{\mbox{\isa{cases}}} to recover this structure in a + sub-proof + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% +\ exE\isanewline +\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% +\ conjE\isanewline +\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% +\ disjE\isanewline +\isanewline +\isacommand{lemma}\isamarkupfalse% +\isanewline +\ \ \isakeyword{assumes}\ A{\isadigit{1}}\ \isakeyword{and}\ A{\isadigit{2}}\ \ % +\isamarkupcmt{assumptions% +} +\isanewline +\ \ \isakeyword{obtains}\isanewline +\ \ \ \ {\isaliteral{28}{\isacharparenleft}}case{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ \ x\ y\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B{\isadigit{1}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}C{\isadigit{1}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{28}{\isacharparenleft}}case{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ \ x\ y\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B{\isadigit{2}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}C{\isadigit{2}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{sorry}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsubsection{Example% +} +\isamarkuptrue% +\isacommand{lemma}\isamarkupfalse% +\ tertium{\isaliteral{5F}{\isacharunderscore}}non{\isaliteral{5F}{\isacharunderscore}}datur{\isaliteral{3A}{\isacharcolon}}\isanewline +\ \ \isakeyword{obtains}\isanewline +\ \ \ \ {\isaliteral{28}{\isacharparenleft}}T{\isaliteral{29}{\isacharparenright}}\ \ A\isanewline +\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{28}{\isacharparenleft}}F{\isaliteral{29}{\isacharparenright}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ x\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ C\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}cases\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ rule{\isaliteral{3A}{\isacharcolon}}\ tertium{\isaliteral{5F}{\isacharunderscore}}non{\isaliteral{5F}{\isacharunderscore}}datur{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ T\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ F\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}x\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ y{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsubsection{Example% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Isabelle/HOL specification mechanisms (datatype, inductive, etc.) + provide suitable derived cases rules.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{datatype}\isamarkupfalse% +\ foo\ {\isaliteral{3D}{\isacharequal}}\ Foo\ {\isaliteral{7C}{\isacharbar}}\ Bar\ foo\isanewline +\isanewline +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ foo\isanewline +\ \ \isacommand{have}\isamarkupfalse% +\ C\isanewline +\ \ \isacommand{proof}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}cases\ x{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ Foo\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ Foo{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{case}\isamarkupfalse% +\ {\isaliteral{28}{\isacharparenleft}}Bar\ a{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ Bar\ a{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% +\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +% +\isamarkupsubsection{Obtaining local contexts% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +A single ``case'' branch may be inlined into Isar proof text + via \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}. This proves \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{22}{\isachardoublequote}}} on the spot, and augments the context afterwards.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{notepad}\isamarkupfalse% +\isanewline +\isakeyword{begin}\isanewline +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isacommand{fix}\isamarkupfalse% +\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline +\isanewline +\ \ \isacommand{obtain}\isamarkupfalse% +\ x\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}B\ x{\isaliteral{60}{\isacharbackquoteclose}}% +\begin{isamarkuptxt}% +Conclusions from this context may not mention \isa{x} again!% +\end{isamarkuptxt}% +\isamarkuptrue% +\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{obtain}\isamarkupfalse% +\ x\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{from}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}B\ x{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{have}\isamarkupfalse% +\ C\ \isacommand{sorry}\isamarkupfalse% +\isanewline +\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% +\isanewline +\ \ \isacommand{note}\isamarkupfalse% +\ {\isaliteral{60}{\isacharbackquoteopen}}C{\isaliteral{60}{\isacharbackquoteclose}}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isacommand{end}\isamarkupfalse% +\isanewline +% +\isadelimtheory +\isanewline +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/IsarRef/isar-ref.tex --- a/doc-src/IsarRef/isar-ref.tex Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/IsarRef/isar-ref.tex Wed Jun 08 10:24:07 2011 +0200 @@ -46,10 +46,13 @@ \maketitle -\pagenumbering{roman} \tableofcontents \clearfirst +\pagenumbering{roman} +{\def\isamarkupchapter#1{\chapter*{#1}}\input{Thy/document/Preface.tex}} +\tableofcontents +\clearfirst \part{Basic Concepts} -\input{Thy/document/Introduction.tex} +\input{Thy/document/Synopsis.tex} \input{Thy/document/Framework.tex} \input{Thy/document/First_Order_Logic.tex} \part{General Language Elements} diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/Ref/Makefile --- a/doc-src/Ref/Makefile Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/Ref/Makefile Wed Jun 08 10:24:07 2011 +0200 @@ -9,10 +9,9 @@ include ../Makefile.in NAME = ref -FILES = ref.tex introduction.tex tactic.tex tctical.tex thm.tex \ - theories.tex defining.tex syntax.tex substitution.tex \ - simplifier.tex classical.tex ../proof.sty ../iman.sty \ - ../extra.sty ../ttbox.sty ../manual.bib +FILES = ref.tex tactic.tex tctical.tex thm.tex defining.tex syntax.tex \ + substitution.tex simplifier.tex classical.tex ../proof.sty \ + ../iman.sty ../extra.sty ../ttbox.sty ../manual.bib dvi: $(NAME).dvi diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/Ref/classical.tex --- a/doc-src/Ref/classical.tex Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/Ref/classical.tex Wed Jun 08 10:24:07 2011 +0200 @@ -3,292 +3,8 @@ \index{classical reasoner|(} \newcommand\ainfer[2]{\begin{array}{r@{\,}l}#2\\ \hline#1\end{array}} -Although Isabelle is generic, many users will be working in some extension of -classical first-order logic. Isabelle's set theory~ZF is built upon -theory~FOL, while HOL conceptually contains first-order logic as a fragment. -Theorem-proving in predicate logic is undecidable, but many researchers have -developed strategies to assist in this task. - -Isabelle's classical reasoner is an \ML{} functor that accepts certain -information about a logic and delivers a suite of automatic tactics. Each -tactic takes a collection of rules and executes a simple, non-clausal proof -procedure. They are slow and simplistic compared with resolution theorem -provers, but they can save considerable time and effort. They can prove -theorems such as Pelletier's~\cite{pelletier86} problems~40 and~41 in -seconds: -\[ (\exists y. \forall x. J(y,x) \bimp \neg J(x,x)) - \imp \neg (\forall x. \exists y. \forall z. J(z,y) \bimp \neg J(z,x)) \] -\[ (\forall z. \exists y. \forall x. F(x,y) \bimp F(x,z) \conj \neg F(x,x)) - \imp \neg (\exists z. \forall x. F(x,z)) -\] -% -The tactics are generic. They are not restricted to first-order logic, and -have been heavily used in the development of Isabelle's set theory. Few -interactive proof assistants provide this much automation. The tactics can -be traced, and their components can be called directly; in this manner, -any proof can be viewed interactively. - -We shall first discuss the underlying principles, then present the classical -reasoner. Finally, we shall see how to instantiate it for new logics. The -logics FOL, ZF, HOL and HOLCF have it already installed. - - -\section{The sequent calculus} -\index{sequent calculus} -Isabelle supports natural deduction, which is easy to use for interactive -proof. But natural deduction does not easily lend itself to automation, -and has a bias towards intuitionism. For certain proofs in classical -logic, it can not be called natural. The {\bf sequent calculus}, a -generalization of natural deduction, is easier to automate. - -A {\bf sequent} has the form $\Gamma\turn\Delta$, where $\Gamma$ -and~$\Delta$ are sets of formulae.% -\footnote{For first-order logic, sequents can equivalently be made from - lists or multisets of formulae.} The sequent -\[ P@1,\ldots,P@m\turn Q@1,\ldots,Q@n \] -is {\bf valid} if $P@1\conj\ldots\conj P@m$ implies $Q@1\disj\ldots\disj -Q@n$. Thus $P@1,\ldots,P@m$ represent assumptions, each of which is true, -while $Q@1,\ldots,Q@n$ represent alternative goals. A sequent is {\bf -basic} if its left and right sides have a common formula, as in $P,Q\turn -Q,R$; basic sequents are trivially valid. - -Sequent rules are classified as {\bf right} or {\bf left}, indicating which -side of the $\turn$~symbol they operate on. Rules that operate on the -right side are analogous to natural deduction's introduction rules, and -left rules are analogous to elimination rules. -Recall the natural deduction rules for - first-order logic, -\iflabelundefined{fol-fig}{from {\it Introduction to Isabelle}}% - {Fig.\ts\ref{fol-fig}}. -The sequent calculus analogue of~$({\imp}I)$ is the rule -$$ -\ainfer{\Gamma &\turn \Delta, P\imp Q}{P,\Gamma &\turn \Delta,Q} -\eqno({\imp}R) -$$ -This breaks down some implication on the right side of a sequent; $\Gamma$ -and $\Delta$ stand for the sets of formulae that are unaffected by the -inference. The analogue of the pair~$({\disj}I1)$ and~$({\disj}I2)$ is the -single rule -$$ -\ainfer{\Gamma &\turn \Delta, P\disj Q}{\Gamma &\turn \Delta,P,Q} -\eqno({\disj}R) -$$ -This breaks down some disjunction on the right side, replacing it by both -disjuncts. Thus, the sequent calculus is a kind of multiple-conclusion logic. - -To illustrate the use of multiple formulae on the right, let us prove -the classical theorem $(P\imp Q)\disj(Q\imp P)$. Working backwards, we -reduce this formula to a basic sequent: -\[ \infer[(\disj)R]{\turn(P\imp Q)\disj(Q\imp P)} - {\infer[(\imp)R]{\turn(P\imp Q), (Q\imp P)\;} - {\infer[(\imp)R]{P \turn Q, (Q\imp P)\qquad} - {P, Q \turn Q, P\qquad\qquad}}} -\] -This example is typical of the sequent calculus: start with the desired -theorem and apply rules backwards in a fairly arbitrary manner. This yields a -surprisingly effective proof procedure. Quantifiers add few complications, -since Isabelle handles parameters and schematic variables. See Chapter~10 -of {\em ML for the Working Programmer}~\cite{paulson-ml2} for further -discussion. - - -\section{Simulating sequents by natural deduction} -Isabelle can represent sequents directly, as in the object-logic~\texttt{LK}\@. -But natural deduction is easier to work with, and most object-logics employ -it. Fortunately, we can simulate the sequent $P@1,\ldots,P@m\turn -Q@1,\ldots,Q@n$ by the Isabelle formula -\[ \List{P@1;\ldots;P@m; \neg Q@2;\ldots; \neg Q@n}\Imp Q@1, \] -where the order of the assumptions and the choice of~$Q@1$ are arbitrary. -Elim-resolution plays a key role in simulating sequent proofs. - -We can easily handle reasoning on the left. -As discussed in -\iflabelundefined{destruct}{{\it Introduction to Isabelle}}{{\S}\ref{destruct}}, -elim-resolution with the rules $(\disj E)$, $(\bot E)$ and $(\exists E)$ -achieves a similar effect as the corresponding sequent rules. For the -other connectives, we use sequent-style elimination rules instead of -destruction rules such as $({\conj}E1,2)$ and $(\forall E)$. But note that -the rule $(\neg L)$ has no effect under our representation of sequents! -$$ -\ainfer{\neg P,\Gamma &\turn \Delta}{\Gamma &\turn \Delta,P}\eqno({\neg}L) -$$ -What about reasoning on the right? Introduction rules can only affect the -formula in the conclusion, namely~$Q@1$. The other right-side formulae are -represented as negated assumptions, $\neg Q@2$, \ldots,~$\neg Q@n$. -\index{assumptions!negated} -In order to operate on one of these, it must first be exchanged with~$Q@1$. -Elim-resolution with the {\bf swap} rule has this effect: -$$ \List{\neg P; \; \neg R\Imp P} \Imp R \eqno(swap) $$ -To ensure that swaps occur only when necessary, each introduction rule is -converted into a swapped form: it is resolved with the second premise -of~$(swap)$. The swapped form of~$({\conj}I)$, which might be -called~$({\neg\conj}E)$, is -\[ \List{\neg(P\conj Q); \; \neg R\Imp P; \; \neg R\Imp Q} \Imp R. \] -Similarly, the swapped form of~$({\imp}I)$ is -\[ \List{\neg(P\imp Q); \; \List{\neg R;P}\Imp Q} \Imp R \] -Swapped introduction rules are applied using elim-resolution, which deletes -the negated formula. Our representation of sequents also requires the use -of ordinary introduction rules. If we had no regard for readability, we -could treat the right side more uniformly by representing sequents as -\[ \List{P@1;\ldots;P@m; \neg Q@1;\ldots; \neg Q@n}\Imp \bot. \] - - -\section{Extra rules for the sequent calculus} -As mentioned, destruction rules such as $({\conj}E1,2)$ and $(\forall E)$ -must be replaced by sequent-style elimination rules. In addition, we need -rules to embody the classical equivalence between $P\imp Q$ and $\neg P\disj -Q$. The introduction rules~$({\disj}I1,2)$ are replaced by a rule that -simulates $({\disj}R)$: -\[ (\neg Q\Imp P) \Imp P\disj Q \] -The destruction rule $({\imp}E)$ is replaced by -\[ \List{P\imp Q;\; \neg P\Imp R;\; Q\Imp R} \Imp R. \] -Quantifier replication also requires special rules. In classical logic, -$\exists x{.}P$ is equivalent to $\neg\forall x{.}\neg P$; the rules -$(\exists R)$ and $(\forall L)$ are dual: -\[ \ainfer{\Gamma &\turn \Delta, \exists x{.}P} - {\Gamma &\turn \Delta, \exists x{.}P, P[t/x]} \; (\exists R) - \qquad - \ainfer{\forall x{.}P, \Gamma &\turn \Delta} - {P[t/x], \forall x{.}P, \Gamma &\turn \Delta} \; (\forall L) -\] -Thus both kinds of quantifier may be replicated. Theorems requiring -multiple uses of a universal formula are easy to invent; consider -\[ (\forall x.P(x)\imp P(f(x))) \conj P(a) \imp P(f^n(a)), \] -for any~$n>1$. Natural examples of the multiple use of an existential -formula are rare; a standard one is $\exists x.\forall y. P(x)\imp P(y)$. - -Forgoing quantifier replication loses completeness, but gains decidability, -since the search space becomes finite. Many useful theorems can be proved -without replication, and the search generally delivers its verdict in a -reasonable time. To adopt this approach, represent the sequent rules -$(\exists R)$, $(\exists L)$ and $(\forall R)$ by $(\exists I)$, $(\exists -E)$ and $(\forall I)$, respectively, and put $(\forall E)$ into elimination -form: -$$ \List{\forall x{.}P(x); P(t)\Imp Q} \Imp Q \eqno(\forall E@2) $$ -Elim-resolution with this rule will delete the universal formula after a -single use. To replicate universal quantifiers, replace the rule by -$$ -\List{\forall x{.}P(x);\; \List{P(t); \forall x{.}P(x)}\Imp Q} \Imp Q. -\eqno(\forall E@3) -$$ -To replicate existential quantifiers, replace $(\exists I)$ by -\[ \List{\neg(\exists x{.}P(x)) \Imp P(t)} \Imp \exists x{.}P(x). \] -All introduction rules mentioned above are also useful in swapped form. - -Replication makes the search space infinite; we must apply the rules with -care. The classical reasoner distinguishes between safe and unsafe -rules, applying the latter only when there is no alternative. Depth-first -search may well go down a blind alley; best-first search is better behaved -in an infinite search space. However, quantifier replication is too -expensive to prove any but the simplest theorems. - - \section{Classical rule sets} \index{classical sets} -Each automatic tactic takes a {\bf classical set} --- a collection of -rules, classified as introduction or elimination and as {\bf safe} or {\bf -unsafe}. In general, safe rules can be attempted blindly, while unsafe -rules must be used with care. A safe rule must never reduce a provable -goal to an unprovable set of subgoals. - -The rule~$({\disj}I1)$ is unsafe because it reduces $P\disj Q$ to~$P$. Any -rule is unsafe whose premises contain new unknowns. The elimination -rule~$(\forall E@2)$ is unsafe, since it is applied via elim-resolution, -which discards the assumption $\forall x{.}P(x)$ and replaces it by the -weaker assumption~$P(\Var{t})$. The rule $({\exists}I)$ is unsafe for -similar reasons. The rule~$(\forall E@3)$ is unsafe in a different sense: -since it keeps the assumption $\forall x{.}P(x)$, it is prone to looping. -In classical first-order logic, all rules are safe except those mentioned -above. - -The safe/unsafe distinction is vague, and may be regarded merely as a way -of giving some rules priority over others. One could argue that -$({\disj}E)$ is unsafe, because repeated application of it could generate -exponentially many subgoals. Induction rules are unsafe because inductive -proofs are difficult to set up automatically. Any inference is unsafe that -instantiates an unknown in the proof state --- thus \ttindex{match_tac} -must be used, rather than \ttindex{resolve_tac}. Even proof by assumption -is unsafe if it instantiates unknowns shared with other subgoals --- thus -\ttindex{eq_assume_tac} must be used, rather than \ttindex{assume_tac}. - -\subsection{Adding rules to classical sets} -Classical rule sets belong to the abstract type \mltydx{claset}, which -supports the following operations (provided the classical reasoner is -installed!): -\begin{ttbox} -empty_cs : claset -print_cs : claset -> unit -rep_cs : claset -> \{safeEs: thm list, safeIs: thm list, - hazEs: thm list, hazIs: thm list, - swrappers: (string * wrapper) list, - uwrappers: (string * wrapper) list, - safe0_netpair: netpair, safep_netpair: netpair, - haz_netpair: netpair, dup_netpair: netpair\} -addSIs : claset * thm list -> claset \hfill{\bf infix 4} -addSEs : claset * thm list -> claset \hfill{\bf infix 4} -addSDs : claset * thm list -> claset \hfill{\bf infix 4} -addIs : claset * thm list -> claset \hfill{\bf infix 4} -addEs : claset * thm list -> claset \hfill{\bf infix 4} -addDs : claset * thm list -> claset \hfill{\bf infix 4} -delrules : claset * thm list -> claset \hfill{\bf infix 4} -\end{ttbox} -The add operations ignore any rule already present in the claset with the same -classification (such as safe introduction). They print a warning if the rule -has already been added with some other classification, but add the rule -anyway. Calling \texttt{delrules} deletes all occurrences of a rule from the -claset, but see the warning below concerning destruction rules. -\begin{ttdescription} -\item[\ttindexbold{empty_cs}] is the empty classical set. - -\item[\ttindexbold{print_cs} $cs$] displays the printable contents of~$cs$, - which is the rules. All other parts are non-printable. - -\item[\ttindexbold{rep_cs} $cs$] decomposes $cs$ as a record of its internal - components, namely the safe introduction and elimination rules, the unsafe - introduction and elimination rules, the lists of safe and unsafe wrappers - (see \ref{sec:modifying-search}), and the internalized forms of the rules. - -\item[$cs$ addSIs $rules$] \indexbold{*addSIs} -adds safe introduction~$rules$ to~$cs$. - -\item[$cs$ addSEs $rules$] \indexbold{*addSEs} -adds safe elimination~$rules$ to~$cs$. - -\item[$cs$ addSDs $rules$] \indexbold{*addSDs} -adds safe destruction~$rules$ to~$cs$. - -\item[$cs$ addIs $rules$] \indexbold{*addIs} -adds unsafe introduction~$rules$ to~$cs$. - -\item[$cs$ addEs $rules$] \indexbold{*addEs} -adds unsafe elimination~$rules$ to~$cs$. - -\item[$cs$ addDs $rules$] \indexbold{*addDs} -adds unsafe destruction~$rules$ to~$cs$. - -\item[$cs$ delrules $rules$] \indexbold{*delrules} -deletes~$rules$ from~$cs$. It prints a warning for those rules that are not -in~$cs$. -\end{ttdescription} - -\begin{warn} - If you added $rule$ using \texttt{addSDs} or \texttt{addDs}, then you must delete - it as follows: -\begin{ttbox} -\(cs\) delrules [make_elim \(rule\)] -\end{ttbox} -\par\noindent -This is necessary because the operators \texttt{addSDs} and \texttt{addDs} convert -the destruction rules to elimination rules by applying \ttindex{make_elim}, -and then insert them using \texttt{addSEs} and \texttt{addEs}, respectively. -\end{warn} - -Introduction rules are those that can be applied using ordinary resolution. -The classical set automatically generates their swapped forms, which will -be applied using elim-resolution. Elimination rules are applied using -elim-resolution. In a classical set, rules are sorted by the number of new -subgoals they will yield; rules that generate the fewest subgoals will be -tried first (see {\S}\ref{biresolve_tac}). For elimination and destruction rules there are variants of the add operations adding a rule in a way such that it is applied only if also its second premise @@ -408,129 +124,28 @@ \section{The classical tactics} -\index{classical reasoner!tactics} If installed, the classical module provides -powerful theorem-proving tactics. - - -\subsection{The tableau prover} -The tactic \texttt{blast_tac} searches for a proof using a fast tableau prover, -coded directly in \ML. It then reconstructs the proof using Isabelle -tactics. It is faster and more powerful than the other classical -reasoning tactics, but has major limitations too. -\begin{itemize} -\item It does not use the wrapper tacticals described above, such as - \ttindex{addss}. -\item It does not perform higher-order unification, as needed by the rule {\tt - rangeI} in HOL and \texttt{RepFunI} in ZF. There are often alternatives - to such rules, for example {\tt range_eqI} and \texttt{RepFun_eqI}. -\item Function variables may only be applied to parameters of the subgoal. -(This restriction arises because the prover does not use higher-order -unification.) If other function variables are present then the prover will -fail with the message {\small\tt Function Var's argument not a bound variable}. -\item Its proof strategy is more general than \texttt{fast_tac}'s but can be - slower. If \texttt{blast_tac} fails or seems to be running forever, try {\tt - fast_tac} and the other tactics described below. -\end{itemize} -% -\begin{ttbox} -blast_tac : claset -> int -> tactic -Blast.depth_tac : claset -> int -> int -> tactic -Blast.trace : bool ref \hfill{\bf initially false} -\end{ttbox} -The two tactics differ on how they bound the number of unsafe steps used in a -proof. While \texttt{blast_tac} starts with a bound of zero and increases it -successively to~20, \texttt{Blast.depth_tac} applies a user-supplied search bound. -\begin{ttdescription} -\item[\ttindexbold{blast_tac} $cs$ $i$] tries to prove - subgoal~$i$, increasing the search bound using iterative - deepening~\cite{korf85}. - -\item[\ttindexbold{Blast.depth_tac} $cs$ $lim$ $i$] tries - to prove subgoal~$i$ using a search bound of $lim$. Sometimes a slow - proof using \texttt{blast_tac} can be made much faster by supplying the - successful search bound to this tactic instead. - -\item[set \ttindexbold{Blast.trace};] \index{tracing!of classical prover} - causes the tableau prover to print a trace of its search. At each step it - displays the formula currently being examined and reports whether the branch - has been closed, extended or split. -\end{ttdescription} - - -\subsection{Automatic tactics}\label{sec:automatic-tactics} -\begin{ttbox} -type clasimpset = claset * simpset; -auto_tac : clasimpset -> tactic -force_tac : clasimpset -> int -> tactic -auto : unit -> unit -force : int -> unit -\end{ttbox} -The automatic tactics attempt to prove goals using a combination of -simplification and classical reasoning. -\begin{ttdescription} -\item[\ttindexbold{auto_tac $(cs,ss)$}] is intended for situations where -there are a lot of mostly trivial subgoals; it proves all the easy ones, -leaving the ones it cannot prove. -(Unfortunately, attempting to prove the hard ones may take a long time.) -\item[\ttindexbold{force_tac} $(cs,ss)$ $i$] is intended to prove subgoal~$i$ -completely. It tries to apply all fancy tactics it knows about, -performing a rather exhaustive search. -\end{ttdescription} - \subsection{Semi-automatic tactics} \begin{ttbox} -clarify_tac : claset -> int -> tactic clarify_step_tac : claset -> int -> tactic -clarsimp_tac : clasimpset -> int -> tactic \end{ttbox} -Use these when the automatic tactics fail. They perform all the obvious -logical inferences that do not split the subgoal. The result is a -simpler subgoal that can be tackled by other means, such as by -instantiating quantifiers yourself. + \begin{ttdescription} -\item[\ttindexbold{clarify_tac} $cs$ $i$] performs a series of safe steps on -subgoal~$i$ by repeatedly calling \texttt{clarify_step_tac}. \item[\ttindexbold{clarify_step_tac} $cs$ $i$] performs a safe step on subgoal~$i$. No splitting step is applied; for example, the subgoal $A\conj B$ is left as a conjunction. Proof by assumption, Modus Ponens, etc., may be performed provided they do not instantiate unknowns. Assumptions of the form $x=t$ may be eliminated. The user-supplied safe wrapper tactical is applied. -\item[\ttindexbold{clarsimp_tac} $cs$ $i$] acts like \texttt{clarify_tac}, but -also does simplification with the given simpset. Note that if the simpset -includes a splitter for the premises, the subgoal may still be split. \end{ttdescription} \subsection{Other classical tactics} \begin{ttbox} -fast_tac : claset -> int -> tactic -best_tac : claset -> int -> tactic -slow_tac : claset -> int -> tactic slow_best_tac : claset -> int -> tactic \end{ttbox} -These tactics attempt to prove a subgoal using sequent-style reasoning. -Unlike \texttt{blast_tac}, they construct proofs directly in Isabelle. Their -effect is restricted (by \texttt{SELECT_GOAL}) to one subgoal; they either prove -this subgoal or fail. The \texttt{slow_} versions conduct a broader -search.% -\footnote{They may, when backtracking from a failed proof attempt, undo even - the step of proving a subgoal by assumption.} -The best-first tactics are guided by a heuristic function: typically, the -total size of the proof state. This function is supplied in the functor call -that sets up the classical reasoner. \begin{ttdescription} -\item[\ttindexbold{fast_tac} $cs$ $i$] applies \texttt{step_tac} using -depth-first search to prove subgoal~$i$. - -\item[\ttindexbold{best_tac} $cs$ $i$] applies \texttt{step_tac} using -best-first search to prove subgoal~$i$. - -\item[\ttindexbold{slow_tac} $cs$ $i$] applies \texttt{slow_step_tac} using -depth-first search to prove subgoal~$i$. - \item[\ttindexbold{slow_best_tac} $cs$ $i$] applies \texttt{slow_step_tac} with best-first search to prove subgoal~$i$. \end{ttdescription} @@ -561,7 +176,6 @@ \subsection{Single-step tactics} \begin{ttbox} safe_step_tac : claset -> int -> tactic -safe_tac : claset -> tactic inst_step_tac : claset -> int -> tactic step_tac : claset -> int -> tactic slow_step_tac : claset -> int -> tactic @@ -574,9 +188,6 @@ include proof by assumption or Modus Ponens (taking care not to instantiate unknowns), or substitution. -\item[\ttindexbold{safe_tac} $cs$] repeatedly performs safe steps on all -subgoals. It is deterministic, with at most one outcome. - \item[\ttindexbold{inst_step_tac} $cs$ $i$] is like \texttt{safe_step_tac}, but allows unknowns to be instantiated. @@ -626,22 +237,6 @@ resolution and also elim-resolution with the swapped form. \end{ttdescription} -\subsection{Creating swapped rules} -\begin{ttbox} -swapify : thm list -> thm list -joinrules : thm list * thm list -> (bool * thm) list -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{swapify} {\it thms}] returns a list consisting of the -swapped versions of~{\it thms}, regarded as introduction rules. - -\item[\ttindexbold{joinrules} ({\it intrs}, {\it elims})] -joins introduction rules, their swapped versions, and elimination rules for -use with \ttindex{biresolve_tac}. Each rule is paired with~\texttt{false} -(indicating ordinary resolution) or~\texttt{true} (indicating -elim-resolution). -\end{ttdescription} - \section{Setting up the classical reasoner}\label{sec:classical-setup} \index{classical reasoner!setting up} diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/Ref/introduction.tex --- a/doc-src/Ref/introduction.tex Wed Jun 08 08:47:43 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,28 +0,0 @@ - -\chapter{Basic Use of Isabelle} - -\section{Ending a session} -\begin{ttbox} -quit : unit -> unit -exit : int -> unit -commit : unit -> bool -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{quit}();] ends the Isabelle session, without saving - the state. - -\item[\ttindexbold{exit} \(i\);] similar to {\tt quit}, passing return - code \(i\) to the operating system. - -\item[\ttindexbold{commit}();] saves the current state without ending - the session, provided that the logic image is opened read-write; - return value {\tt false} indicates an error. -\end{ttdescription} - -Typing control-D also finishes the session in essentially the same way -as the sequence {\tt commit(); quit();} would. - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "ref" -%%% End: diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/Ref/ref.tex --- a/doc-src/Ref/ref.tex Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/Ref/ref.tex Wed Jun 08 10:24:07 2011 +0200 @@ -36,11 +36,9 @@ \subsubsection*{Acknowledgements} Tobias Nipkow, of T. U. Munich, wrote most of - Chapters~\protect\ref{Defining-Logics} and~\protect\ref{chap:simplification}, - and part of - Chapter~\protect\ref{theories}. Carsten Clasohm also contributed to - Chapter~\protect\ref{theories}. Markus Wenzel contributed to - Chapter~\protect\ref{chap:syntax}. Jeremy Dawson, Sara Kalvala, Martin + Chapters~\protect\ref{Defining-Logics} and~\protect\ref{chap:simplification}. + Markus Wenzel contributed to Chapter~\protect\ref{chap:syntax}. + Jeremy Dawson, Sara Kalvala, Martin Simons and others suggested changes and corrections. The research has been funded by the EPSRC (grants GR/G53279, GR/H40570, GR/K57381, GR/K77051, GR/M75440) and by ESPRIT @@ -49,11 +47,9 @@ \pagenumbering{roman} \tableofcontents \clearfirst -\include{introduction} \include{tactic} \include{tctical} \include{thm} -\include{theories} \include{defining} \include{syntax} \include{substitution} diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/Ref/simplifier.tex --- a/doc-src/Ref/simplifier.tex Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/Ref/simplifier.tex Wed Jun 08 10:24:07 2011 +0200 @@ -336,47 +336,6 @@ \index{rewrite rules|)} -\subsection{*Simplification procedures} -\begin{ttbox} -addsimprocs : simpset * simproc list -> simpset -delsimprocs : simpset * simproc list -> simpset -\end{ttbox} - -Simplification procedures are {\ML} objects of abstract type -\texttt{simproc}. Basically they are just functions that may produce -\emph{proven} rewrite rules on demand. They are associated with -certain patterns that conceptually represent left-hand sides of -equations; these are shown by \texttt{print_ss}. During its -operation, the simplifier may offer a simplification procedure the -current redex and ask for a suitable rewrite rule. Thus rules may be -specifically fashioned for particular situations, resulting in a more -powerful mechanism than term rewriting by a fixed set of rules. - - -\begin{ttdescription} - -\item[$ss$ \ttindexbold{addsimprocs} $procs$] adds the simplification - procedures $procs$ to the current simpset. - -\item[$ss$ \ttindexbold{delsimprocs} $procs$] deletes the simplification - procedures $procs$ from the current simpset. - -\end{ttdescription} - -For example, simplification procedures \ttindexbold{nat_cancel} of -\texttt{HOL/Arith} cancel common summands and constant factors out of -several relations of sums over natural numbers. - -Consider the following goal, which after cancelling $a$ on both sides -contains a factor of $2$. Simplifying with the simpset of -\texttt{Arith.thy} will do the cancellation automatically: -\begin{ttbox} -{\out 1. x + a + x < y + y + 2 + a + a + a + a + a} -by (Simp_tac 1); -{\out 1. x < Suc (a + (a + y))} -\end{ttbox} - - \subsection{*Congruence rules}\index{congruence rules}\label{sec:simp-congs} \begin{ttbox} addcongs : simpset * thm list -> simpset \hfill{\bf infix 4} @@ -920,82 +879,6 @@ \index{rewrite rules!permutative|)} -\section{*Coding simplification procedures} -\begin{ttbox} - val Simplifier.simproc: Sign.sg -> string -> string list - -> (Sign.sg -> simpset -> term -> thm option) -> simproc - val Simplifier.simproc_i: Sign.sg -> string -> term list - -> (Sign.sg -> simpset -> term -> thm option) -> simproc -\end{ttbox} - -\begin{ttdescription} -\item[\ttindexbold{Simplifier.simproc}~$sign$~$name$~$lhss$~$proc$] makes - $proc$ a simplification procedure for left-hand side patterns $lhss$. The - name just serves as a comment. The function $proc$ may be invoked by the - simplifier for redex positions matched by one of $lhss$ as described below - (which are be specified as strings to be read as terms). - -\item[\ttindexbold{Simplifier.simproc_i}] is similar to - \verb,Simplifier.simproc,, but takes well-typed terms as pattern argument. -\end{ttdescription} - -Simplification procedures are applied in a two-stage process as -follows: The simplifier tries to match the current redex position -against any one of the $lhs$ patterns of any simplification procedure. -If this succeeds, it invokes the corresponding {\ML} function, passing -with the current signature, local assumptions and the (potential) -redex. The result may be either \texttt{None} (indicating failure) or -\texttt{Some~$thm$}. - -Any successful result is supposed to be a (possibly conditional) -rewrite rule $t \equiv u$ that is applicable to the current redex. -The rule will be applied just as any ordinary rewrite rule. It is -expected to be already in \emph{internal form}, though, bypassing the -automatic preprocessing of object-level equivalences. - -\medskip - -As an example of how to write your own simplification procedures, -consider eta-expansion of pair abstraction (see also -\texttt{HOL/Modelcheck/MCSyn} where this is used to provide external -model checker syntax). - -The HOL theory of tuples (see \texttt{HOL/Prod}) provides an operator -\texttt{split} together with some concrete syntax supporting -$\lambda\,(x,y).b$ abstractions. Assume that we would like to offer a tactic -that rewrites any function $\lambda\,p.f\,p$ (where $p$ is of some pair type) -to $\lambda\,(x,y).f\,(x,y)$. The corresponding rule is: -\begin{ttbox} -pair_eta_expand: (f::'a*'b=>'c) = (\%(x, y). f (x, y)) -\end{ttbox} -Unfortunately, term rewriting using this rule directly would not -terminate! We now use the simplification procedure mechanism in order -to stop the simplifier from applying this rule over and over again, -making it rewrite only actual abstractions. The simplification -procedure \texttt{pair_eta_expand_proc} is defined as follows: -\begin{ttbox} -val pair_eta_expand_proc = - Simplifier.simproc (Theory.sign_of (the_context ())) - "pair_eta_expand" ["f::'a*'b=>'c"] - (fn _ => fn _ => fn t => - case t of Abs _ => Some (mk_meta_eq pair_eta_expand) - | _ => None); -\end{ttbox} -This is an example of using \texttt{pair_eta_expand_proc}: -\begin{ttbox} -{\out 1. P (\%p::'a * 'a. fst p + snd p + z)} -by (simp_tac (simpset() addsimprocs [pair_eta_expand_proc]) 1); -{\out 1. P (\%(x::'a,y::'a). x + y + z)} -\end{ttbox} - -\medskip - -In the above example the simplification procedure just did fine -grained control over rule application, beyond higher-order pattern -matching. Usually, procedures would do some more work, in particular -prove particular theorems depending on the current redex. - - \section{*Setting up the Simplifier}\label{sec:setting-up-simp} \index{simplification!setting up} diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/Ref/theories.tex --- a/doc-src/Ref/theories.tex Wed Jun 08 08:47:43 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,287 +0,0 @@ - -\chapter{Theories, Terms and Types} \label{theories} -\index{theories|(} - -\section{Basic operations on theories}\label{BasicOperationsOnTheories} - -\subsection{*Theory inclusion} -\begin{ttbox} -transfer : theory -> thm -> thm -\end{ttbox} - -Transferring theorems to super theories has no logical significance, -but may affect some operations in subtle ways (e.g.\ implicit merges -of signatures when applying rules, or pretty printing of theorems). - -\begin{ttdescription} - -\item[\ttindexbold{transfer} $thy$ $thm$] transfers theorem $thm$ to - theory $thy$, provided the latter includes the theory of $thm$. - -\end{ttdescription} - - -\section{Terms}\label{sec:terms} -\index{terms|bold} -Terms belong to the \ML\ type \mltydx{term}, which is a concrete datatype -with six constructors: -\begin{ttbox} -type indexname = string * int; -infix 9 $; -datatype term = Const of string * typ - | Free of string * typ - | Var of indexname * typ - | Bound of int - | Abs of string * typ * term - | op $ of term * term; -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{Const} ($a$, $T$)] \index{constants|bold} - is the \textbf{constant} with name~$a$ and type~$T$. Constants include - connectives like $\land$ and $\forall$ as well as constants like~0 - and~$Suc$. Other constants may be required to define a logic's concrete - syntax. - -\item[\ttindexbold{Free} ($a$, $T$)] \index{variables!free|bold} - is the \textbf{free variable} with name~$a$ and type~$T$. - -\item[\ttindexbold{Var} ($v$, $T$)] \index{unknowns|bold} - is the \textbf{scheme variable} with indexname~$v$ and type~$T$. An - \mltydx{indexname} is a string paired with a non-negative index, or - subscript; a term's scheme variables can be systematically renamed by - incrementing their subscripts. Scheme variables are essentially free - variables, but may be instantiated during unification. - -\item[\ttindexbold{Bound} $i$] \index{variables!bound|bold} - is the \textbf{bound variable} with de Bruijn index~$i$, which counts the - number of lambdas, starting from zero, between a variable's occurrence - and its binding. The representation prevents capture of variables. For - more information see de Bruijn \cite{debruijn72} or - Paulson~\cite[page~376]{paulson-ml2}. - -\item[\ttindexbold{Abs} ($a$, $T$, $u$)] - \index{lambda abs@$\lambda$-abstractions|bold} - is the $\lambda$-\textbf{abstraction} with body~$u$, and whose bound - variable has name~$a$ and type~$T$. The name is used only for parsing - and printing; it has no logical significance. - -\item[$t$ \$ $u$] \index{$@{\tt\$}|bold} \index{function applications|bold} -is the \textbf{application} of~$t$ to~$u$. -\end{ttdescription} -Application is written as an infix operator to aid readability. Here is an -\ML\ pattern to recognize FOL formulae of the form~$A\imp B$, binding the -subformulae to~$A$ and~$B$: -\begin{ttbox} -Const("Trueprop",_) $ (Const("op -->",_) $ A $ B) -\end{ttbox} - - -\section{*Variable binding} -\begin{ttbox} -loose_bnos : term -> int list -incr_boundvars : int -> term -> term -abstract_over : term*term -> term -variant_abs : string * typ * term -> string * term -aconv : term * term -> bool\hfill\textbf{infix} -\end{ttbox} -These functions are all concerned with the de Bruijn representation of -bound variables. -\begin{ttdescription} -\item[\ttindexbold{loose_bnos} $t$] - returns the list of all dangling bound variable references. In - particular, \texttt{Bound~0} is loose unless it is enclosed in an - abstraction. Similarly \texttt{Bound~1} is loose unless it is enclosed in - at least two abstractions; if enclosed in just one, the list will contain - the number 0. A well-formed term does not contain any loose variables. - -\item[\ttindexbold{incr_boundvars} $j$] - increases a term's dangling bound variables by the offset~$j$. This is - required when moving a subterm into a context where it is enclosed by a - different number of abstractions. Bound variables with a matching - abstraction are unaffected. - -\item[\ttindexbold{abstract_over} $(v,t)$] - forms the abstraction of~$t$ over~$v$, which may be any well-formed term. - It replaces every occurrence of \(v\) by a \texttt{Bound} variable with the - correct index. - -\item[\ttindexbold{variant_abs} $(a,T,u)$] - substitutes into $u$, which should be the body of an abstraction. - It replaces each occurrence of the outermost bound variable by a free - variable. The free variable has type~$T$ and its name is a variant - of~$a$ chosen to be distinct from all constants and from all variables - free in~$u$. - -\item[$t$ \ttindexbold{aconv} $u$] - tests whether terms~$t$ and~$u$ are \(\alpha\)-convertible: identical up - to renaming of bound variables. -\begin{itemize} - \item - Two constants, \texttt{Free}s, or \texttt{Var}s are \(\alpha\)-convertible - if their names and types are equal. - (Variables having the same name but different types are thus distinct. - This confusing situation should be avoided!) - \item - Two bound variables are \(\alpha\)-convertible - if they have the same number. - \item - Two abstractions are \(\alpha\)-convertible - if their bodies are, and their bound variables have the same type. - \item - Two applications are \(\alpha\)-convertible - if the corresponding subterms are. -\end{itemize} - -\end{ttdescription} - -\section{Certified terms}\index{terms!certified|bold}\index{signatures} -A term $t$ can be \textbf{certified} under a signature to ensure that every type -in~$t$ is well-formed and every constant in~$t$ is a type instance of a -constant declared in the signature. The term must be well-typed and its use -of bound variables must be well-formed. Meta-rules such as \texttt{forall_elim} -take certified terms as arguments. - -Certified terms belong to the abstract type \mltydx{cterm}. -Elements of the type can only be created through the certification process. -In case of error, Isabelle raises exception~\ttindex{TERM}\@. - -\subsection{Printing terms} -\index{terms!printing of} -\begin{ttbox} - string_of_cterm : cterm -> string -Sign.string_of_term : Sign.sg -> term -> string -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{string_of_cterm} $ct$] -displays $ct$ as a string. - -\item[\ttindexbold{Sign.string_of_term} $sign$ $t$] -displays $t$ as a string, using the syntax of~$sign$. -\end{ttdescription} - -\subsection{Making and inspecting certified terms} -\begin{ttbox} -cterm_of : Sign.sg -> term -> cterm -read_cterm : Sign.sg -> string * typ -> cterm -cert_axm : Sign.sg -> string * term -> string * term -read_axm : Sign.sg -> string * string -> string * term -rep_cterm : cterm -> \{T:typ, t:term, sign:Sign.sg, maxidx:int\} -Sign.certify_term : Sign.sg -> term -> term * typ * int -\end{ttbox} -\begin{ttdescription} - -\item[\ttindexbold{cterm_of} $sign$ $t$] \index{signatures} certifies - $t$ with respect to signature~$sign$. - -\item[\ttindexbold{read_cterm} $sign$ ($s$, $T$)] reads the string~$s$ - using the syntax of~$sign$, creating a certified term. The term is - checked to have type~$T$; this type also tells the parser what kind - of phrase to parse. - -\item[\ttindexbold{cert_axm} $sign$ ($name$, $t$)] certifies $t$ with - respect to $sign$ as a meta-proposition and converts all exceptions - to an error, including the final message -\begin{ttbox} -The error(s) above occurred in axiom "\(name\)" -\end{ttbox} - -\item[\ttindexbold{read_axm} $sign$ ($name$, $s$)] similar to {\tt - cert_axm}, but first reads the string $s$ using the syntax of - $sign$. - -\item[\ttindexbold{rep_cterm} $ct$] decomposes $ct$ as a record - containing its type, the term itself, its signature, and the maximum - subscript of its unknowns. The type and maximum subscript are - computed during certification. - -\item[\ttindexbold{Sign.certify_term}] is a more primitive version of - \texttt{cterm_of}, returning the internal representation instead of - an abstract \texttt{cterm}. - -\end{ttdescription} - - -\section{Types}\index{types|bold} -Types belong to the \ML\ type \mltydx{typ}, which is a concrete datatype with -three constructor functions. These correspond to type constructors, free -type variables and schematic type variables. Types are classified by sorts, -which are lists of classes (representing an intersection). A class is -represented by a string. -\begin{ttbox} -type class = string; -type sort = class list; - -datatype typ = Type of string * typ list - | TFree of string * sort - | TVar of indexname * sort; - -infixr 5 -->; -fun S --> T = Type ("fun", [S, T]); -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{Type} ($a$, $Ts$)] \index{type constructors|bold} - applies the \textbf{type constructor} named~$a$ to the type operand list~$Ts$. - Type constructors include~\tydx{fun}, the binary function space - constructor, as well as nullary type constructors such as~\tydx{prop}. - Other type constructors may be introduced. In expressions, but not in - patterns, \hbox{\tt$S$-->$T$} is a convenient shorthand for function - types. - -\item[\ttindexbold{TFree} ($a$, $s$)] \index{type variables|bold} - is the \textbf{type variable} with name~$a$ and sort~$s$. - -\item[\ttindexbold{TVar} ($v$, $s$)] \index{type unknowns|bold} - is the \textbf{type unknown} with indexname~$v$ and sort~$s$. - Type unknowns are essentially free type variables, but may be - instantiated during unification. -\end{ttdescription} - - -\section{Certified types} -\index{types!certified|bold} -Certified types, which are analogous to certified terms, have type -\ttindexbold{ctyp}. - -\subsection{Printing types} -\index{types!printing of} -\begin{ttbox} - string_of_ctyp : ctyp -> string -Sign.string_of_typ : Sign.sg -> typ -> string -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{string_of_ctyp} $cT$] -displays $cT$ as a string. - -\item[\ttindexbold{Sign.string_of_typ} $sign$ $T$] -displays $T$ as a string, using the syntax of~$sign$. -\end{ttdescription} - - -\subsection{Making and inspecting certified types} -\begin{ttbox} -ctyp_of : Sign.sg -> typ -> ctyp -rep_ctyp : ctyp -> \{T: typ, sign: Sign.sg\} -Sign.certify_typ : Sign.sg -> typ -> typ -\end{ttbox} -\begin{ttdescription} - -\item[\ttindexbold{ctyp_of} $sign$ $T$] \index{signatures} certifies - $T$ with respect to signature~$sign$. - -\item[\ttindexbold{rep_ctyp} $cT$] decomposes $cT$ as a record - containing the type itself and its signature. - -\item[\ttindexbold{Sign.certify_typ}] is a more primitive version of - \texttt{ctyp_of}, returning the internal representation instead of - an abstract \texttt{ctyp}. - -\end{ttdescription} - - -\index{theories|)} - - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "ref" -%%% End: diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/Ref/thm.tex --- a/doc-src/Ref/thm.tex Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/Ref/thm.tex Wed Jun 08 10:24:07 2011 +0200 @@ -3,58 +3,13 @@ \index{theorems|(} Theorems, which represent the axioms, theorems and rules of -object-logics, have type \mltydx{thm}. This chapter begins by -describing operations that print theorems and that join them in -forward proof. Most theorem operations are intended for advanced -applications, such as programming new proof procedures. Many of these -operations refer to signatures, certified terms and certified types, -which have the \ML{} types {\tt Sign.sg}, {\tt cterm} and {\tt ctyp} -and are discussed in Chapter~\ref{theories}. Beginning users should -ignore such complexities --- and skip all but the first section of -this chapter. +object-logics, have type \mltydx{thm}. This chapter describes +operations that join theorems in forward proof. Most theorem +operations are intended for advanced applications, such as programming +new proof procedures. \section{Basic operations on theorems} -\subsection{Pretty-printing a theorem} -\index{theorems!printing of} -\begin{ttbox} -prth : thm -> thm -prths : thm list -> thm list -prthq : thm Seq.seq -> thm Seq.seq -print_thm : thm -> unit -print_goals : int -> thm -> unit -string_of_thm : thm -> string -\end{ttbox} -The first three commands are for interactive use. They are identity -functions that display, then return, their argument. The \ML{} identifier -{\tt it} will refer to the value just displayed. - -The others are for use in programs. Functions with result type {\tt unit} -are convenient for imperative programming. - -\begin{ttdescription} -\item[\ttindexbold{prth} {\it thm}] -prints {\it thm\/} at the terminal. - -\item[\ttindexbold{prths} {\it thms}] -prints {\it thms}, a list of theorems. - -\item[\ttindexbold{prthq} {\it thmq}] -prints {\it thmq}, a sequence of theorems. It is useful for inspecting -the output of a tactic. - -\item[\ttindexbold{print_thm} {\it thm}] -prints {\it thm\/} at the terminal. - -\item[\ttindexbold{print_goals} {\it limit\/} {\it thm}] -prints {\it thm\/} in goal style, with the premises as subgoals. It prints -at most {\it limit\/} subgoals. The subgoal module calls {\tt print_goals} -to display proof states. - -\item[\ttindexbold{string_of_thm} {\it thm}] -converts {\it thm\/} to a string. -\end{ttdescription} - \subsection{Forward proof: joining rules by resolution} \index{theorems!joining by resolution} @@ -371,123 +326,6 @@ \section{*Primitive meta-level inference rules} \index{meta-rules|(} -These implement the meta-logic in the style of the {\sc lcf} system, -as functions from theorems to theorems. They are, rarely, useful for -deriving results in the pure theory. Mainly, they are included for -completeness, and most users should not bother with them. The -meta-rules raise exception \xdx{THM} to signal malformed premises, -incompatible signatures and similar errors. - -\index{meta-assumptions} -The meta-logic uses natural deduction. Each theorem may depend on -meta-level assumptions. Certain rules, such as $({\Imp}I)$, -discharge assumptions; in most other rules, the conclusion depends on all -of the assumptions of the premises. Formally, the system works with -assertions of the form -\[ \phi \quad [\phi@1,\ldots,\phi@n], \] -where $\phi@1$,~\ldots,~$\phi@n$ are the assumptions. This can be -also read as a single conclusion sequent $\phi@1,\ldots,\phi@n \vdash -\phi$. Do not confuse meta-level assumptions with the object-level -assumptions in a subgoal, which are represented in the meta-logic -using~$\Imp$. - -Each theorem has a signature. Certified terms have a signature. When a -rule takes several premises and certified terms, it merges the signatures -to make a signature for the conclusion. This fails if the signatures are -incompatible. - -\medskip - -The following presentation of primitive rules ignores sort -hypotheses\index{sort hypotheses} (see also \S\ref{sec:sort-hyps}). These are -handled transparently by the logic implementation. - -\bigskip - -\index{meta-implication} -The \textbf{implication} rules are $({\Imp}I)$ -and $({\Imp}E)$: -\[ \infer[({\Imp}I)]{\phi\Imp \psi}{\infer*{\psi}{[\phi]}} \qquad - \infer[({\Imp}E)]{\psi}{\phi\Imp \psi & \phi} \] - -\index{meta-equality} -Equality of truth values means logical equivalence: -\[ \infer[({\equiv}I)]{\phi\equiv\psi}{\phi\Imp\psi & - \psi\Imp\phi} - \qquad - \infer[({\equiv}E)]{\psi}{\phi\equiv \psi & \phi} \] - -The \textbf{equality} rules are reflexivity, symmetry, and transitivity: -\[ {a\equiv a}\,(refl) \qquad - \infer[(sym)]{b\equiv a}{a\equiv b} \qquad - \infer[(trans)]{a\equiv c}{a\equiv b & b\equiv c} \] - -\index{lambda calc@$\lambda$-calculus} -The $\lambda$-conversions are $\alpha$-conversion, $\beta$-conversion, and -extensionality:\footnote{$\alpha$-conversion holds if $y$ is not free -in~$a$; $(ext)$ holds if $x$ is not free in the assumptions, $f$, or~$g$.} -\[ {(\lambda x.a) \equiv (\lambda y.a[y/x])} \qquad - {((\lambda x.a)(b)) \equiv a[b/x]} \qquad - \infer[(ext)]{f\equiv g}{f(x) \equiv g(x)} \] - -The \textbf{abstraction} and \textbf{combination} rules let conversions be -applied to subterms:\footnote{Abstraction holds if $x$ is not free in the -assumptions.} -\[ \infer[(abs)]{(\lambda x.a) \equiv (\lambda x.b)}{a\equiv b} \qquad - \infer[(comb)]{f(a)\equiv g(b)}{f\equiv g & a\equiv b} \] - -\index{meta-quantifiers} -The \textbf{universal quantification} rules are $(\Forall I)$ and $(\Forall -E)$:\footnote{$(\Forall I)$ holds if $x$ is not free in the assumptions.} -\[ \infer[(\Forall I)]{\Forall x.\phi}{\phi} \qquad - \infer[(\Forall E)]{\phi[b/x]}{\Forall x.\phi} \] - - -\subsection{Assumption rule} -\index{meta-assumptions} -\begin{ttbox} -assume: cterm -> thm -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{assume} $ct$] -makes the theorem \(\phi \;[\phi]\), where $\phi$ is the value of~$ct$. -The rule checks that $ct$ has type $prop$ and contains no unknowns, which -are not allowed in assumptions. -\end{ttdescription} - -\subsection{Implication rules} -\index{meta-implication} -\begin{ttbox} -implies_intr : cterm -> thm -> thm -implies_intr_list : cterm list -> thm -> thm -implies_intr_hyps : thm -> thm -implies_elim : thm -> thm -> thm -implies_elim_list : thm -> thm list -> thm -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{implies_intr} $ct$ $thm$] -is $({\Imp}I)$, where $ct$ is the assumption to discharge, say~$\phi$. It -maps the premise~$\psi$ to the conclusion $\phi\Imp\psi$, removing all -occurrences of~$\phi$ from the assumptions. The rule checks that $ct$ has -type $prop$. - -\item[\ttindexbold{implies_intr_list} $cts$ $thm$] -applies $({\Imp}I)$ repeatedly, on every element of the list~$cts$. - -\item[\ttindexbold{implies_intr_hyps} $thm$] -applies $({\Imp}I)$ to discharge all the hypotheses (assumptions) of~$thm$. -It maps the premise $\phi \; [\phi@1,\ldots,\phi@n]$ to the conclusion -$\List{\phi@1,\ldots,\phi@n}\Imp\phi$. - -\item[\ttindexbold{implies_elim} $thm@1$ $thm@2$] -applies $({\Imp}E)$ to $thm@1$ and~$thm@2$. It maps the premises $\phi\Imp -\psi$ and $\phi$ to the conclusion~$\psi$. - -\item[\ttindexbold{implies_elim_list} $thm$ $thms$] -applies $({\Imp}E)$ repeatedly to $thm$, using each element of~$thms$ in -turn. It maps the premises $\List{\phi@1,\ldots,\phi@n}\Imp\psi$ and -$\phi@1$,\ldots,$\phi@n$ to the conclusion~$\psi$. -\end{ttdescription} \subsection{Logical equivalence rules} \index{meta-equality} @@ -561,97 +399,6 @@ \end{ttdescription} -\subsection{Forall introduction rules} -\index{meta-quantifiers} -\begin{ttbox} -forall_intr : cterm -> thm -> thm -forall_intr_list : cterm list -> thm -> thm -forall_intr_frees : thm -> thm -\end{ttbox} - -\begin{ttdescription} -\item[\ttindexbold{forall_intr} $x$ $thm$] -applies $({\Forall}I)$, abstracting over all occurrences (if any!) of~$x$. -The rule maps the premise $\phi$ to the conclusion $\Forall x.\phi$. -Parameter~$x$ is supplied as a cterm. It may be an unknown or a free -variable (provided it does not occur in the assumptions). - -\item[\ttindexbold{forall_intr_list} $xs$ $thm$] -applies $({\Forall}I)$ repeatedly, on every element of the list~$xs$. - -\item[\ttindexbold{forall_intr_frees} $thm$] -applies $({\Forall}I)$ repeatedly, generalizing over all the free variables -of the premise. -\end{ttdescription} - - -\subsection{Forall elimination rules} -\begin{ttbox} -forall_elim : cterm -> thm -> thm -forall_elim_list : cterm list -> thm -> thm -forall_elim_var : int -> thm -> thm -forall_elim_vars : int -> thm -> thm -\end{ttbox} - -\begin{ttdescription} -\item[\ttindexbold{forall_elim} $ct$ $thm$] -applies $({\Forall}E)$, mapping the premise $\Forall x.\phi$ to the conclusion -$\phi[ct/x]$. The rule checks that $ct$ and $x$ have the same type. - -\item[\ttindexbold{forall_elim_list} $cts$ $thm$] -applies $({\Forall}E)$ repeatedly, on every element of the list~$cts$. - -\item[\ttindexbold{forall_elim_var} $k$ $thm$] -applies $({\Forall}E)$, mapping the premise $\Forall x.\phi$ to the conclusion -$\phi[\Var{x@k}/x]$. Thus, it replaces the outermost $\Forall$-bound -variable by an unknown having subscript~$k$. - -\item[\ttindexbold{forall_elim_vars} $k$ $thm$] -applies {\tt forall_elim_var}~$k$ repeatedly until the theorem no longer has -the form $\Forall x.\phi$. -\end{ttdescription} - - -\subsection{Instantiation of unknowns} -\index{instantiation} -\begin{alltt}\footnotesize -instantiate: (indexname * ctyp){\thinspace}list * (cterm * cterm){\thinspace}list -> thm -> thm -\end{alltt} -There are two versions of this rule. The primitive one is -\ttindexbold{Thm.instantiate}, which merely performs the instantiation and can -produce a conclusion not in normal form. A derived version is -\ttindexbold{Drule.instantiate}, which normalizes its conclusion. - -\begin{ttdescription} -\item[\ttindexbold{instantiate} ($tyinsts$,$insts$) $thm$] -simultaneously substitutes types for type unknowns (the -$tyinsts$) and terms for term unknowns (the $insts$). Instantiations are -given as $(v,t)$ pairs, where $v$ is an unknown and $t$ is a term (of the -same type as $v$) or a type (of the same sort as~$v$). All the unknowns -must be distinct. - -In some cases, \ttindex{instantiate'} (see \S\ref{sec:instantiate}) -provides a more convenient interface to this rule. -\end{ttdescription} - - - - -\subsection{Freezing/thawing type unknowns} -\index{type unknowns!freezing/thawing of} -\begin{ttbox} -freezeT: thm -> thm -varifyT: thm -> thm -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{freezeT} $thm$] -converts all the type unknowns in $thm$ to free type variables. - -\item[\ttindexbold{varifyT} $thm$] -converts all the free type variables in $thm$ to type unknowns. -\end{ttdescription} - - \section{Derived rules for goal-directed proof} Most of these rules have the sole purpose of implementing particular tactics. There are few occasions for applying them directly to a theorem. diff -r 3535f16d9714 -r bc72c1ccc89e doc-src/manual.bib --- a/doc-src/manual.bib Wed Jun 08 08:47:43 2011 +0200 +++ b/doc-src/manual.bib Wed Jun 08 10:24:07 2011 +0200 @@ -556,8 +556,7 @@ @Book{mgordon-hol, editor = {M. J. C. Gordon and T. F. Melham}, - title = {Introduction to {HOL}: A Theorem Proving Environment for - Higher Order Logic}, + title = {Introduction to {HOL}: A Theorem Proving Environment for Higher Order Logic}, publisher = CUP, year = 1993} @@ -1343,6 +1342,15 @@ year = 1986, note = {Errata, JAR 4 (1988), 235--236 and JAR 18 (1997), 135}} +@InCollection{pitts93, + author = {A. Pitts}, + title = {The {HOL} Logic}, + editor = {M. J. C. Gordon and T. F. Melham}, + booktitle = {Introduction to {HOL}: A Theorem Proving Environment for Higher Order Logic}, + pages = {191--232}, + publisher = CUP, + year = 1993} + @Article{pitts94, author = {Andrew M. Pitts}, title = {A Co-induction Principle for Recursively Defined Domains},