| author | haftmann | 
| Sun, 14 Jun 2009 09:11:51 +0200 | |
| changeset 31627 | bc2de3795756 | 
| parent 27167 | a99747ccba87 | 
| child 32836 | 4c6e3e7ac2bf | 
| permissions | -rw-r--r-- | 
| 10365 | 1 | % | 
| 11187 | 2 | \begin{isabellebody}%
 | 
| 3 | \def\isabellecontext{Advanced}%
 | |
| 17056 | 4 | % | 
| 5 | \isadelimtheory | |
| 6 | % | |
| 7 | \endisadelimtheory | |
| 8 | % | |
| 9 | \isatagtheory | |
| 23848 | 10 | % | 
| 17056 | 11 | \endisatagtheory | 
| 12 | {\isafoldtheory}%
 | |
| 13 | % | |
| 14 | \isadelimtheory | |
| 15 | % | |
| 16 | \endisadelimtheory | |
| 23848 | 17 | % | 
| 18 | \begin{isamarkuptext}%
 | |
| 19 | The premises of introduction rules may contain universal quantifiers and | |
| 20 | monotone functions. A universal quantifier lets the rule | |
| 21 | refer to any number of instances of | |
| 22 | the inductively defined set. A monotone function lets the rule refer | |
| 23 | to existing constructions (such as ``list of'') over the inductively defined | |
| 24 | set. The examples below show how to use the additional expressiveness | |
| 25 | and how to reason from the resulting definitions.% | |
| 26 | \end{isamarkuptext}%
 | |
| 27 | \isamarkuptrue% | |
| 28 | % | |
| 29 | \isamarkupsubsection{Universal Quantifiers in Introduction Rules \label{sec:gterm-datatype}%
 | |
| 30 | } | |
| 31 | \isamarkuptrue% | |
| 32 | % | |
| 33 | \begin{isamarkuptext}%
 | |
| 34 | \index{ground terms example|(}%
 | |
| 35 | \index{quantifiers!and inductive definitions|(}%
 | |
| 36 | As a running example, this section develops the theory of \textbf{ground
 | |
| 37 | terms}: terms constructed from constant and function | |
| 38 | symbols but not variables. To simplify matters further, we regard a | |
| 39 | constant as a function applied to the null argument list. Let us declare a | |
| 40 | datatype \isa{gterm} for the type of ground  terms. It is a type constructor
 | |
| 41 | whose argument is a type of function symbols.% | |
| 42 | \end{isamarkuptext}%
 | |
| 43 | \isamarkuptrue% | |
| 17175 | 44 | \isacommand{datatype}\isamarkupfalse%
 | 
| 23848 | 45 | \ {\isacharprime}f\ gterm\ {\isacharequal}\ Apply\ {\isacharprime}f\ {\isachardoublequoteopen}{\isacharprime}f\ gterm\ list{\isachardoublequoteclose}%
 | 
| 46 | \begin{isamarkuptext}%
 | |
| 47 | To try it out, we declare a datatype of some integer operations: | |
| 48 | integer constants, the unary minus operator and the addition | |
| 49 | operator.% | |
| 50 | \end{isamarkuptext}%
 | |
| 51 | \isamarkuptrue% | |
| 17175 | 52 | \isacommand{datatype}\isamarkupfalse%
 | 
| 23848 | 53 | \ integer{\isacharunderscore}op\ {\isacharequal}\ Number\ int\ {\isacharbar}\ UnaryMinus\ {\isacharbar}\ Plus%
 | 
| 54 | \begin{isamarkuptext}%
 | |
| 55 | Now the type \isa{integer{\isacharunderscore}op\ gterm} denotes the ground 
 | |
| 56 | terms built over those symbols. | |
| 57 | ||
| 58 | The type constructor \isa{gterm} can be generalized to a function 
 | |
| 59 | over sets. It returns | |
| 60 | the set of ground terms that can be formed over a set \isa{F} of function symbols. For
 | |
| 61 | example, we could consider the set of ground terms formed from the finite | |
| 62 | set \isa{{\isacharbraceleft}Number\ {\isadigit{2}}{\isacharcomma}\ UnaryMinus{\isacharcomma}\ Plus{\isacharbraceright}}.
 | |
| 63 | ||
| 64 | This concept is inductive. If we have a list \isa{args} of ground terms 
 | |
| 65 | over~\isa{F} and a function symbol \isa{f} in \isa{F}, then we 
 | |
| 66 | can apply \isa{f} to \isa{args} to obtain another ground term. 
 | |
| 67 | The only difficulty is that the argument list may be of any length. Hitherto, | |
| 68 | each rule in an inductive definition referred to the inductively | |
| 69 | defined set a fixed number of times, typically once or twice. | |
| 70 | A universal quantifier in the premise of the introduction rule | |
| 71 | expresses that every element of \isa{args} belongs
 | |
| 72 | to our inductively defined set: is a ground term | |
| 73 | over~\isa{F}.  The function \isa{set} denotes the set of elements in a given 
 | |
| 74 | list.% | |
| 75 | \end{isamarkuptext}%
 | |
| 76 | \isamarkuptrue% | |
| 23733 | 77 | \isacommand{inductive{\isacharunderscore}set}\isamarkupfalse%
 | 
| 78 | \isanewline | |
| 79 | \ \ gterms\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}f\ set\ {\isasymRightarrow}\ {\isacharprime}f\ gterm\ set{\isachardoublequoteclose}\isanewline
 | |
| 80 | \ \ \isakeyword{for}\ F\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}f\ set{\isachardoublequoteclose}\isanewline
 | |
| 81 | \isakeyword{where}\isanewline
 | |
| 17175 | 82 | step{\isacharbrackleft}intro{\isacharbang}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}{\isasymforall}t\ {\isasymin}\ set\ args{\isachardot}\ t\ {\isasymin}\ gterms\ F{\isacharsemicolon}\ \ f\ {\isasymin}\ F{\isasymrbrakk}\isanewline
 | 
| 23848 | 83 | \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isacharparenleft}Apply\ f\ args{\isacharparenright}\ {\isasymin}\ gterms\ F{\isachardoublequoteclose}%
 | 
| 84 | \begin{isamarkuptext}%
 | |
| 85 | To demonstrate a proof from this definition, let us | |
| 86 | show that the function \isa{gterms}
 | |
| 87 | is \textbf{monotone}.  We shall need this concept shortly.%
 | |
| 88 | \end{isamarkuptext}%
 | |
| 89 | \isamarkuptrue% | |
| 17175 | 90 | \isacommand{lemma}\isamarkupfalse%
 | 
| 91 | \ gterms{\isacharunderscore}mono{\isacharcolon}\ {\isachardoublequoteopen}F{\isasymsubseteq}G\ {\isasymLongrightarrow}\ gterms\ F\ {\isasymsubseteq}\ gterms\ G{\isachardoublequoteclose}\isanewline
 | |
| 17056 | 92 | % | 
| 93 | \isadelimproof | |
| 94 | % | |
| 95 | \endisadelimproof | |
| 96 | % | |
| 97 | \isatagproof | |
| 17175 | 98 | \isacommand{apply}\isamarkupfalse%
 | 
| 99 | \ clarify\isanewline | |
| 100 | \isacommand{apply}\isamarkupfalse%
 | |
| 23848 | 101 | \ {\isacharparenleft}erule\ gterms{\isachardot}induct{\isacharparenright}\isanewline
 | 
| 17175 | 102 | \isacommand{apply}\isamarkupfalse%
 | 
| 103 | \ blast\isanewline | |
| 104 | \isacommand{done}\isamarkupfalse%
 | |
| 105 | % | |
| 17056 | 106 | \endisatagproof | 
| 107 | {\isafoldproof}%
 | |
| 108 | % | |
| 109 | \isadelimproof | |
| 110 | % | |
| 111 | \endisadelimproof | |
| 11866 | 112 | % | 
| 23848 | 113 | \isadelimproof | 
| 114 | % | |
| 115 | \endisadelimproof | |
| 116 | % | |
| 117 | \isatagproof | |
| 118 | % | |
| 119 | \begin{isamarkuptxt}%
 | |
| 120 | Intuitively, this theorem says that | |
| 121 | enlarging the set of function symbols enlarges the set of ground | |
| 122 | terms. The proof is a trivial rule induction. | |
| 123 | First we use the \isa{clarify} method to assume the existence of an element of
 | |
| 124 | \isa{gterms\ F}.  (We could have used \isa{intro\ subsetI}.)  We then
 | |
| 125 | apply rule induction. Here is the resulting subgoal: | |
| 11187 | 126 | \begin{isabelle}%
 | 
| 23848 | 127 | \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ args\ f{\isachardot}\isanewline
 | 
| 128 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymlbrakk}F\ {\isasymsubseteq}\ G{\isacharsemicolon}\ {\isasymforall}t{\isasymin}set\ args{\isachardot}\ t\ {\isasymin}\ gterms\ F\ {\isasymand}\ t\ {\isasymin}\ gterms\ G{\isacharsemicolon}\ f\ {\isasymin}\ F{\isasymrbrakk}\isanewline
 | |
| 129 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ Apply\ f\ args\ {\isasymin}\ gterms\ G%
 | |
| 11173 | 130 | \end{isabelle}
 | 
| 23848 | 131 | The assumptions state that \isa{f} belongs 
 | 
| 132 | to~\isa{F}, which is included in~\isa{G}, and that every element of the list \isa{args} is
 | |
| 133 | a ground term over~\isa{G}.  The \isa{blast} method finds this chain of reasoning easily.%
 | |
| 134 | \end{isamarkuptxt}%
 | |
| 135 | \isamarkuptrue% | |
| 136 | % | |
| 137 | \endisatagproof | |
| 138 | {\isafoldproof}%
 | |
| 139 | % | |
| 140 | \isadelimproof | |
| 141 | % | |
| 142 | \endisadelimproof | |
| 143 | % | |
| 144 | \begin{isamarkuptext}%
 | |
| 145 | \begin{warn}
 | |
| 146 | Why do we call this function \isa{gterms} instead 
 | |
| 147 | of \isa{gterm}?  A constant may have the same name as a type.  However,
 | |
| 148 | name clashes could arise in the theorems that Isabelle generates. | |
| 149 | Our choice of names keeps \isa{gterms{\isachardot}induct} separate from 
 | |
| 150 | \isa{gterm{\isachardot}induct}.
 | |
| 151 | \end{warn}
 | |
| 10469 | 152 | |
| 23848 | 153 | Call a term \textbf{well-formed} if each symbol occurring in it is applied
 | 
| 154 | to the correct number of arguments. (This number is called the symbol's | |
| 155 | \textbf{arity}.)  We can express well-formedness by
 | |
| 156 | generalizing the inductive definition of | |
| 157 | \isa{gterms}.
 | |
| 158 | Suppose we are given a function called \isa{arity}, specifying the arities
 | |
| 159 | of all symbols.  In the inductive step, we have a list \isa{args} of such
 | |
| 160 | terms and a function  symbol~\isa{f}. If the length of the list matches the
 | |
| 161 | function's arity  then applying \isa{f} to \isa{args} yields a well-formed
 | |
| 162 | term.% | |
| 163 | \end{isamarkuptext}%
 | |
| 164 | \isamarkuptrue% | |
| 165 | \isacommand{inductive{\isacharunderscore}set}\isamarkupfalse%
 | |
| 166 | \isanewline | |
| 167 | \ \ well{\isacharunderscore}formed{\isacharunderscore}gterm\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}f\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}f\ gterm\ set{\isachardoublequoteclose}\isanewline
 | |
| 168 | \ \ \isakeyword{for}\ arity\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}f\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline
 | |
| 169 | \isakeyword{where}\isanewline
 | |
| 170 | step{\isacharbrackleft}intro{\isacharbang}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}{\isasymforall}t\ {\isasymin}\ set\ args{\isachardot}\ t\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity{\isacharsemicolon}\ \ \isanewline
 | |
| 171 | \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ length\ args\ {\isacharequal}\ arity\ f{\isasymrbrakk}\isanewline
 | |
| 172 | \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isacharparenleft}Apply\ f\ args{\isacharparenright}\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity{\isachardoublequoteclose}%
 | |
| 173 | \begin{isamarkuptext}%
 | |
| 174 | The inductive definition neatly captures the reasoning above. | |
| 175 | The universal quantification over the | |
| 176 | \isa{set} of arguments expresses that all of them are well-formed.%
 | |
| 177 | \index{quantifiers!and inductive definitions|)}%
 | |
| 11187 | 178 | \end{isamarkuptext}%
 | 
| 17175 | 179 | \isamarkuptrue% | 
| 23848 | 180 | % | 
| 181 | \isamarkupsubsection{Alternative Definition Using a Monotone Function%
 | |
| 182 | } | |
| 183 | \isamarkuptrue% | |
| 184 | % | |
| 185 | \begin{isamarkuptext}%
 | |
| 186 | \index{monotone functions!and inductive definitions|(}% 
 | |
| 187 | An inductive definition may refer to the | |
| 188 | inductively defined set through an arbitrary monotone function. To | |
| 189 | demonstrate this powerful feature, let us | |
| 190 | change the inductive definition above, replacing the | |
| 191 | quantifier by a use of the function \isa{lists}. This
 | |
| 192 | function, from the Isabelle theory of lists, is analogous to the | |
| 193 | function \isa{gterms} declared above: if \isa{A} is a set then
 | |
| 194 | \isa{lists\ A} is the set of lists whose elements belong to
 | |
| 195 | \isa{A}.  
 | |
| 196 | ||
| 197 | In the inductive definition of well-formed terms, examine the one | |
| 198 | introduction rule.  The first premise states that \isa{args} belongs to
 | |
| 199 | the \isa{lists} of well-formed terms.  This formulation is more
 | |
| 200 | direct, if more obscure, than using a universal quantifier.% | |
| 201 | \end{isamarkuptext}%
 | |
| 202 | \isamarkuptrue% | |
| 203 | \isacommand{inductive{\isacharunderscore}set}\isamarkupfalse%
 | |
| 10365 | 204 | \isanewline | 
| 23848 | 205 | \ \ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}f\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}f\ gterm\ set{\isachardoublequoteclose}\isanewline
 | 
| 206 | \ \ \isakeyword{for}\ arity\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}f\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline
 | |
| 207 | \isakeyword{where}\isanewline
 | |
| 208 | step{\isacharbrackleft}intro{\isacharbang}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymlbrakk}args\ {\isasymin}\ lists\ {\isacharparenleft}well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity{\isacharparenright}{\isacharsemicolon}\ \ \isanewline
 | |
| 209 | \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ length\ args\ {\isacharequal}\ arity\ f{\isasymrbrakk}\isanewline
 | |
| 210 | \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isacharparenleft}Apply\ f\ args{\isacharparenright}\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity{\isachardoublequoteclose}\isanewline
 | |
| 211 | \isakeyword{monos}\ lists{\isacharunderscore}mono%
 | |
| 11187 | 212 | \begin{isamarkuptext}%
 | 
| 23848 | 213 | We cite the theorem \isa{lists{\isacharunderscore}mono} to justify 
 | 
| 214 | using the function \isa{lists}.%
 | |
| 215 | \footnote{This particular theorem is installed by default already, but we
 | |
| 216 | include the \isakeyword{monos} declaration in order to illustrate its syntax.}
 | |
| 11187 | 217 | \begin{isabelle}%
 | 
| 23848 | 218 | A\ {\isasymsubseteq}\ B\ {\isasymLongrightarrow}\ lists\ A\ {\isasymsubseteq}\ lists\ B\rulename{lists{\isacharunderscore}mono}%
 | 
| 219 | \end{isabelle}
 | |
| 220 | Why must the function be monotone? An inductive definition describes | |
| 221 | an iterative construction: each element of the set is constructed by a | |
| 222 | finite number of introduction rule applications. For example, the | |
| 223 | elements of \isa{even} are constructed by finitely many applications of
 | |
| 224 | the rules | |
| 225 | \begin{isabelle}%
 | |
| 226 | {\isadigit{0}}\ {\isasymin}\ even\isasep\isanewline%
 | |
| 227 | n\ {\isasymin}\ even\ {\isasymLongrightarrow}\ Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isasymin}\ even%
 | |
| 11187 | 228 | \end{isabelle}
 | 
| 23848 | 229 | All references to a set in its | 
| 230 | inductive definition must be positive. Applications of an | |
| 231 | introduction rule cannot invalidate previous applications, allowing the | |
| 232 | construction process to converge. | |
| 233 | The following pair of rules do not constitute an inductive definition: | |
| 234 | \begin{trivlist}
 | |
| 235 | \item \isa{{\isadigit{0}}\ {\isasymin}\ even}
 | |
| 236 | \item \isa{n\ {\isasymnotin}\ even\ {\isasymLongrightarrow}\ Suc\ n\ {\isasymin}\ even}
 | |
| 237 | \end{trivlist}
 | |
| 238 | Showing that 4 is even using these rules requires showing that 3 is not | |
| 239 | even. It is far from trivial to show that this set of rules | |
| 240 | characterizes the even numbers. | |
| 10469 | 241 | |
| 23848 | 242 | Even with its use of the function \isa{lists}, the premise of our
 | 
| 243 | introduction rule is positive: | |
| 244 | \begin{isabelle}%
 | |
| 245 | args\ {\isasymin}\ lists\ {\isacharparenleft}well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity{\isacharparenright}%
 | |
| 246 | \end{isabelle}
 | |
| 247 | To apply the rule we construct a list \isa{args} of previously
 | |
| 248 | constructed well-formed terms. We obtain a | |
| 249 | new term, \isa{Apply\ f\ args}.  Because \isa{lists} is monotone,
 | |
| 250 | applications of the rule remain valid as new terms are constructed. | |
| 251 | Further lists of well-formed | |
| 252 | terms become available and none are taken away.% | |
| 253 | \index{monotone functions!and inductive definitions|)}%
 | |
| 254 | \end{isamarkuptext}%
 | |
| 255 | \isamarkuptrue% | |
| 256 | % | |
| 257 | \isamarkupsubsection{A Proof of Equivalence%
 | |
| 258 | } | |
| 259 | \isamarkuptrue% | |
| 260 | % | |
| 261 | \begin{isamarkuptext}%
 | |
| 262 | We naturally hope that these two inductive definitions of ``well-formed'' | |
| 263 | coincide. The equality can be proved by separate inclusions in | |
| 264 | each direction. Each is a trivial rule induction.% | |
| 11187 | 265 | \end{isamarkuptext}%
 | 
| 17175 | 266 | \isamarkuptrue% | 
| 267 | \isacommand{lemma}\isamarkupfalse%
 | |
| 23848 | 268 | \ {\isachardoublequoteopen}well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity\ {\isasymsubseteq}\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity{\isachardoublequoteclose}\isanewline
 | 
| 17056 | 269 | % | 
| 270 | \isadelimproof | |
| 271 | % | |
| 272 | \endisadelimproof | |
| 273 | % | |
| 274 | \isatagproof | |
| 17175 | 275 | \isacommand{apply}\isamarkupfalse%
 | 
| 23848 | 276 | \ clarify\isanewline | 
| 277 | \isacommand{apply}\isamarkupfalse%
 | |
| 278 | \ {\isacharparenleft}erule\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isachardot}induct{\isacharparenright}\isanewline
 | |
| 279 | \isacommand{apply}\isamarkupfalse%
 | |
| 280 | \ auto\isanewline | |
| 281 | \isacommand{done}\isamarkupfalse%
 | |
| 282 | % | |
| 283 | \endisatagproof | |
| 284 | {\isafoldproof}%
 | |
| 285 | % | |
| 286 | \isadelimproof | |
| 287 | % | |
| 288 | \endisadelimproof | |
| 289 | % | |
| 290 | \isadelimproof | |
| 291 | % | |
| 292 | \endisadelimproof | |
| 293 | % | |
| 294 | \isatagproof | |
| 295 | % | |
| 296 | \begin{isamarkuptxt}%
 | |
| 297 | The \isa{clarify} method gives
 | |
| 298 | us an element of \isa{well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity} on which to perform 
 | |
| 299 | induction. The resulting subgoal can be proved automatically: | |
| 300 | \begin{isabelle}%
 | |
| 301 | \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ args\ f{\isachardot}\isanewline
 | |
| 302 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymlbrakk}{\isasymforall}t{\isasymin}set\ args{\isachardot}\isanewline
 | |
| 303 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymlbrakk}\ \ \ }t\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity\ {\isasymand}\ t\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity{\isacharsemicolon}\isanewline
 | |
| 304 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ \ }length\ args\ {\isacharequal}\ arity\ f{\isasymrbrakk}\isanewline
 | |
| 305 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ Apply\ f\ args\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity%
 | |
| 306 | \end{isabelle}
 | |
| 307 | This proof resembles the one given in | |
| 308 | {\S}\ref{sec:gterm-datatype} above, especially in the form of the
 | |
| 309 | induction hypothesis. Next, we consider the opposite inclusion:% | |
| 310 | \end{isamarkuptxt}%
 | |
| 311 | \isamarkuptrue% | |
| 312 | % | |
| 313 | \endisatagproof | |
| 314 | {\isafoldproof}%
 | |
| 315 | % | |
| 316 | \isadelimproof | |
| 317 | % | |
| 318 | \endisadelimproof | |
| 319 | \isacommand{lemma}\isamarkupfalse%
 | |
| 320 | \ {\isachardoublequoteopen}well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity\ {\isasymsubseteq}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity{\isachardoublequoteclose}\isanewline
 | |
| 321 | % | |
| 322 | \isadelimproof | |
| 323 | % | |
| 324 | \endisadelimproof | |
| 325 | % | |
| 326 | \isatagproof | |
| 327 | \isacommand{apply}\isamarkupfalse%
 | |
| 328 | \ clarify\isanewline | |
| 329 | \isacommand{apply}\isamarkupfalse%
 | |
| 330 | \ {\isacharparenleft}erule\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}{\isachardot}induct{\isacharparenright}\isanewline
 | |
| 331 | \isacommand{apply}\isamarkupfalse%
 | |
| 332 | \ auto\isanewline | |
| 333 | \isacommand{done}\isamarkupfalse%
 | |
| 17175 | 334 | % | 
| 17056 | 335 | \endisatagproof | 
| 336 | {\isafoldproof}%
 | |
| 337 | % | |
| 338 | \isadelimproof | |
| 23848 | 339 | % | 
| 340 | \endisadelimproof | |
| 341 | % | |
| 342 | \isadelimproof | |
| 17056 | 343 | % | 
| 344 | \endisadelimproof | |
| 23848 | 345 | % | 
| 346 | \isatagproof | |
| 347 | % | |
| 348 | \begin{isamarkuptxt}%
 | |
| 27167 | 349 | The proof script is virtually identical, | 
| 350 | but the subgoal after applying induction may be surprising: | |
| 23848 | 351 | \begin{isabelle}%
 | 
| 352 | \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ args\ f{\isachardot}\isanewline
 | |
| 353 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymlbrakk}args\isanewline
 | |
| 354 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymlbrakk}}{\isasymin}\ lists\isanewline
 | |
| 355 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymlbrakk}{\isasymin}\ \ }{\isacharparenleft}well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity\ {\isasyminter}\isanewline
 | |
| 356 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymlbrakk}{\isasymin}\ \ {\isacharparenleft}}{\isacharbraceleft}a{\isachardot}\ a\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity{\isacharbraceright}{\isacharparenright}{\isacharsemicolon}\isanewline
 | |
| 357 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ \ }length\ args\ {\isacharequal}\ arity\ f{\isasymrbrakk}\isanewline
 | |
| 358 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ Apply\ f\ args\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity%
 | |
| 359 | \end{isabelle}
 | |
| 360 | The induction hypothesis contains an application of \isa{lists}.  Using a
 | |
| 361 | monotone function in the inductive definition always has this effect. The | |
| 362 | subgoal may look uninviting, but fortunately | |
| 363 | \isa{lists} distributes over intersection:
 | |
| 364 | \begin{isabelle}%
 | |
| 365 | lists\ {\isacharparenleft}A\ {\isasyminter}\ B{\isacharparenright}\ {\isacharequal}\ lists\ A\ {\isasyminter}\ lists\ B\rulename{lists{\isacharunderscore}Int{\isacharunderscore}eq}%
 | |
| 366 | \end{isabelle}
 | |
| 367 | Thanks to this default simplification rule, the induction hypothesis | |
| 368 | is quickly replaced by its two parts: | |
| 369 | \begin{trivlist}
 | |
| 370 | \item \isa{args\ {\isasymin}\ lists\ {\isacharparenleft}well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity{\isacharparenright}}
 | |
| 371 | \item \isa{args\ {\isasymin}\ lists\ {\isacharparenleft}well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity{\isacharparenright}}
 | |
| 372 | \end{trivlist}
 | |
| 373 | Invoking the rule \isa{well{\isacharunderscore}formed{\isacharunderscore}gterm{\isachardot}step} completes the proof.  The
 | |
| 374 | call to \isa{auto} does all this work.
 | |
| 375 | ||
| 376 | This example is typical of how monotone functions | |
| 377 | \index{monotone functions} can be used.  In particular, many of them
 | |
| 378 | distribute over intersection. Monotonicity implies one direction of | |
| 379 | this set equality; we have this theorem: | |
| 380 | \begin{isabelle}%
 | |
| 381 | mono\ f\ {\isasymLongrightarrow}\ f\ {\isacharparenleft}A\ {\isasyminter}\ B{\isacharparenright}\ {\isasymsubseteq}\ f\ A\ {\isasyminter}\ f\ B\rulename{mono{\isacharunderscore}Int}%
 | |
| 382 | \end{isabelle}%
 | |
| 383 | \end{isamarkuptxt}%
 | |
| 384 | \isamarkuptrue% | |
| 385 | % | |
| 386 | \endisatagproof | |
| 387 | {\isafoldproof}%
 | |
| 388 | % | |
| 389 | \isadelimproof | |
| 390 | % | |
| 391 | \endisadelimproof | |
| 392 | % | |
| 393 | \isamarkupsubsection{Another Example of Rule Inversion%
 | |
| 394 | } | |
| 395 | \isamarkuptrue% | |
| 396 | % | |
| 397 | \begin{isamarkuptext}%
 | |
| 398 | \index{rule inversion|(}%
 | |
| 399 | Does \isa{gterms} distribute over intersection?  We have proved that this
 | |
| 400 | function is monotone, so \isa{mono{\isacharunderscore}Int} gives one of the inclusions.  The
 | |
| 401 | opposite inclusion asserts that if \isa{t} is a ground term over both of the
 | |
| 402 | sets | |
| 403 | \isa{F} and~\isa{G} then it is also a ground term over their intersection,
 | |
| 404 | \isa{F\ {\isasyminter}\ G}.%
 | |
| 405 | \end{isamarkuptext}%
 | |
| 406 | \isamarkuptrue% | |
| 407 | \isacommand{lemma}\isamarkupfalse%
 | |
| 408 | \ gterms{\isacharunderscore}IntI{\isacharcolon}\isanewline
 | |
| 409 | \ \ \ \ \ {\isachardoublequoteopen}t\ {\isasymin}\ gterms\ F\ {\isasymLongrightarrow}\ t\ {\isasymin}\ gterms\ G\ {\isasymlongrightarrow}\ t\ {\isasymin}\ gterms\ {\isacharparenleft}F{\isasyminter}G{\isacharparenright}{\isachardoublequoteclose}%
 | |
| 410 | \isadelimproof | |
| 411 | % | |
| 412 | \endisadelimproof | |
| 413 | % | |
| 414 | \isatagproof | |
| 415 | % | |
| 416 | \endisatagproof | |
| 417 | {\isafoldproof}%
 | |
| 418 | % | |
| 419 | \isadelimproof | |
| 420 | % | |
| 421 | \endisadelimproof | |
| 422 | % | |
| 423 | \begin{isamarkuptext}%
 | |
| 424 | Attempting this proof, we get the assumption | |
| 425 | \isa{Apply\ f\ args\ {\isasymin}\ gterms\ G}, which cannot be broken down. 
 | |
| 426 | It looks like a job for rule inversion:\cmmdx{inductive\protect\_cases}%
 | |
| 427 | \end{isamarkuptext}%
 | |
| 428 | \isamarkuptrue% | |
| 17175 | 429 | \isacommand{inductive{\isacharunderscore}cases}\isamarkupfalse%
 | 
| 430 | \ gterm{\isacharunderscore}Apply{\isacharunderscore}elim\ {\isacharbrackleft}elim{\isacharbang}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}Apply\ f\ args\ {\isasymin}\ gterms\ F{\isachardoublequoteclose}%
 | |
| 11187 | 431 | \begin{isamarkuptext}%
 | 
| 23848 | 432 | Here is the result. | 
| 11187 | 433 | \begin{isabelle}%
 | 
| 23848 | 434 | {\isasymlbrakk}Apply\ f\ args\ {\isasymin}\ gterms\ F{\isacharsemicolon}\isanewline
 | 
| 435 | \isaindent{\ }{\isasymlbrakk}{\isasymforall}t{\isasymin}set\ args{\isachardot}\ t\ {\isasymin}\ gterms\ F{\isacharsemicolon}\ f\ {\isasymin}\ F{\isasymrbrakk}\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\isanewline
 | |
| 436 | {\isasymLongrightarrow}\ P\rulename{gterm{\isacharunderscore}Apply{\isacharunderscore}elim}%
 | |
| 10469 | 437 | \end{isabelle}
 | 
| 23848 | 438 | This rule replaces an assumption about \isa{Apply\ f\ args} by 
 | 
| 439 | assumptions about \isa{f} and~\isa{args}.  
 | |
| 440 | No cases are discarded (there was only one to begin | |
| 441 | with) but the rule applies specifically to the pattern \isa{Apply\ f\ args}.
 | |
| 442 | It can be applied repeatedly as an elimination rule without looping, so we | |
| 443 | have given the \isa{elim{\isacharbang}} attribute. 
 | |
| 444 | ||
| 445 | Now we can prove the other half of that distributive law.% | |
| 11187 | 446 | \end{isamarkuptext}%
 | 
| 17175 | 447 | \isamarkuptrue% | 
| 448 | \isacommand{lemma}\isamarkupfalse%
 | |
| 449 | \ gterms{\isacharunderscore}IntI\ {\isacharbrackleft}rule{\isacharunderscore}format{\isacharcomma}\ intro{\isacharbang}{\isacharbrackright}{\isacharcolon}\isanewline
 | |
| 450 | \ \ \ \ \ {\isachardoublequoteopen}t\ {\isasymin}\ gterms\ F\ {\isasymLongrightarrow}\ t\ {\isasymin}\ gterms\ G\ {\isasymlongrightarrow}\ t\ {\isasymin}\ gterms\ {\isacharparenleft}F{\isasyminter}G{\isacharparenright}{\isachardoublequoteclose}\isanewline
 | |
| 17056 | 451 | % | 
| 452 | \isadelimproof | |
| 453 | % | |
| 454 | \endisadelimproof | |
| 455 | % | |
| 456 | \isatagproof | |
| 17175 | 457 | \isacommand{apply}\isamarkupfalse%
 | 
| 23848 | 458 | \ {\isacharparenleft}erule\ gterms{\isachardot}induct{\isacharparenright}\isanewline
 | 
| 17175 | 459 | \isacommand{apply}\isamarkupfalse%
 | 
| 460 | \ blast\isanewline | |
| 461 | \isacommand{done}\isamarkupfalse%
 | |
| 462 | % | |
| 17056 | 463 | \endisatagproof | 
| 464 | {\isafoldproof}%
 | |
| 465 | % | |
| 466 | \isadelimproof | |
| 467 | % | |
| 468 | \endisadelimproof | |
| 11866 | 469 | % | 
| 23848 | 470 | \isadelimproof | 
| 471 | % | |
| 472 | \endisadelimproof | |
| 473 | % | |
| 474 | \isatagproof | |
| 475 | % | |
| 476 | \begin{isamarkuptxt}%
 | |
| 477 | The proof begins with rule induction over the definition of | |
| 478 | \isa{gterms}, which leaves a single subgoal:  
 | |
| 11187 | 479 | \begin{isabelle}%
 | 
| 23848 | 480 | \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}args\ f{\isachardot}\isanewline
 | 
| 481 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymlbrakk}{\isasymforall}t{\isasymin}set\ args{\isachardot}\isanewline
 | |
| 482 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymlbrakk}\ \ \ }t\ {\isasymin}\ gterms\ F\ {\isasymand}\ {\isacharparenleft}t\ {\isasymin}\ gterms\ G\ {\isasymlongrightarrow}\ t\ {\isasymin}\ gterms\ {\isacharparenleft}F\ {\isasyminter}\ G{\isacharparenright}{\isacharparenright}{\isacharsemicolon}\isanewline
 | |
| 483 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ \ }f\ {\isasymin}\ F{\isasymrbrakk}\isanewline
 | |
| 484 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ Apply\ f\ args\ {\isasymin}\ gterms\ G\ {\isasymlongrightarrow}\isanewline
 | |
| 485 | \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymLongrightarrow}\ }Apply\ f\ args\ {\isasymin}\ gterms\ {\isacharparenleft}F\ {\isasyminter}\ G{\isacharparenright}%
 | |
| 11173 | 486 | \end{isabelle}
 | 
| 23848 | 487 | To prove this, we assume \isa{Apply\ f\ args\ {\isasymin}\ gterms\ G}.  Rule inversion,
 | 
| 488 | in the form of \isa{gterm{\isacharunderscore}Apply{\isacharunderscore}elim}, infers
 | |
| 489 | that every element of \isa{args} belongs to 
 | |
| 490 | \isa{gterms\ G}; hence (by the induction hypothesis) it belongs
 | |
| 491 | to \isa{gterms\ {\isacharparenleft}F\ {\isasyminter}\ G{\isacharparenright}}.  Rule inversion also yields
 | |
| 492 | \isa{f\ {\isasymin}\ G} and hence \isa{f\ {\isasymin}\ F\ {\isasyminter}\ G}. 
 | |
| 493 | All of this reasoning is done by \isa{blast}.
 | |
| 494 | ||
| 495 | \smallskip | |
| 496 | Our distributive law is a trivial consequence of previously-proved results:% | |
| 497 | \end{isamarkuptxt}%
 | |
| 17175 | 498 | \isamarkuptrue% | 
| 23848 | 499 | % | 
| 500 | \endisatagproof | |
| 501 | {\isafoldproof}%
 | |
| 502 | % | |
| 503 | \isadelimproof | |
| 504 | % | |
| 505 | \endisadelimproof | |
| 17175 | 506 | \isacommand{lemma}\isamarkupfalse%
 | 
| 507 | \ gterms{\isacharunderscore}Int{\isacharunderscore}eq\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\isanewline
 | |
| 23848 | 508 | \ \ \ \ \ {\isachardoublequoteopen}gterms\ {\isacharparenleft}F\ {\isasyminter}\ G{\isacharparenright}\ {\isacharequal}\ gterms\ F\ {\isasyminter}\ gterms\ G{\isachardoublequoteclose}\isanewline
 | 
| 17056 | 509 | % | 
| 510 | \isadelimproof | |
| 511 | % | |
| 512 | \endisadelimproof | |
| 513 | % | |
| 514 | \isatagproof | |
| 17175 | 515 | \isacommand{by}\isamarkupfalse%
 | 
| 516 | \ {\isacharparenleft}blast\ intro{\isacharbang}{\isacharcolon}\ mono{\isacharunderscore}Int\ monoI\ gterms{\isacharunderscore}mono{\isacharparenright}%
 | |
| 17056 | 517 | \endisatagproof | 
| 518 | {\isafoldproof}%
 | |
| 519 | % | |
| 520 | \isadelimproof | |
| 521 | % | |
| 522 | \endisadelimproof | |
| 12156 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 paulson parents: 
11866diff
changeset | 523 | % | 
| 23848 | 524 | \index{rule inversion|)}%
 | 
| 525 | \index{ground terms example|)}
 | |
| 526 | ||
| 527 | ||
| 528 | \begin{isamarkuptext}
 | |
| 529 | \begin{exercise}
 | |
| 530 | A function mapping function symbols to their | |
| 531 | types is called a \textbf{signature}.  Given a type 
 | |
| 532 | ranging over type symbols, we can represent a function's type by a | |
| 533 | list of argument types paired with the result type. | |
| 534 | Complete this inductive definition: | |
| 535 | \begin{isabelle}
 | |
| 23733 | 536 | \isacommand{inductive{\isacharunderscore}set}\isamarkupfalse%
 | 
| 537 | \isanewline | |
| 23848 | 538 | \ \ well{\isacharunderscore}typed{\isacharunderscore}gterm\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}f\ {\isasymRightarrow}\ {\isacharprime}t\ list\ {\isacharasterisk}\ {\isacharprime}t{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}f\ gterm\ {\isacharasterisk}\ {\isacharprime}t{\isacharparenright}set{\isachardoublequoteclose}\isanewline
 | 
| 539 | \ \ \isakeyword{for}\ sig\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}f\ {\isasymRightarrow}\ {\isacharprime}t\ list\ {\isacharasterisk}\ {\isacharprime}t{\isachardoublequoteclose}%
 | |
| 540 | \end{isabelle}
 | |
| 541 | \end{exercise}
 | |
| 542 | \end{isamarkuptext}
 | |
| 17056 | 543 | % | 
| 544 | \isadelimproof | |
| 545 | % | |
| 546 | \endisadelimproof | |
| 547 | % | |
| 548 | \isatagproof | |
| 17175 | 549 | % | 
| 17056 | 550 | \endisatagproof | 
| 551 | {\isafoldproof}%
 | |
| 552 | % | |
| 553 | \isadelimproof | |
| 554 | % | |
| 555 | \endisadelimproof | |
| 11866 | 556 | % | 
| 17056 | 557 | \isadelimproof | 
| 558 | % | |
| 559 | \endisadelimproof | |
| 560 | % | |
| 561 | \isatagproof | |
| 17175 | 562 | % | 
| 17056 | 563 | \endisatagproof | 
| 564 | {\isafoldproof}%
 | |
| 565 | % | |
| 566 | \isadelimproof | |
| 567 | % | |
| 568 | \endisadelimproof | |
| 569 | % | |
| 570 | \isadelimtheory | |
| 571 | % | |
| 572 | \endisadelimtheory | |
| 573 | % | |
| 574 | \isatagtheory | |
| 17175 | 575 | % | 
| 17056 | 576 | \endisatagtheory | 
| 577 | {\isafoldtheory}%
 | |
| 578 | % | |
| 579 | \isadelimtheory | |
| 580 | % | |
| 581 | \endisadelimtheory | |
| 11187 | 582 | \end{isabellebody}%
 | 
| 10365 | 583 | %%% Local Variables: | 
| 584 | %%% mode: latex | |
| 585 | %%% TeX-master: "root" | |
| 586 | %%% End: |