| 10879 |      1 | % $Id$
 | 
|  |      2 | The premises of introduction rules may contain universal quantifiers and
 | 
| 11173 |      3 | monotone functions.  A universal quantifier lets the rule 
 | 
|  |      4 | refer to any number of instances of 
 | 
|  |      5 | the inductively defined set.  A monotone function lets the rule refer
 | 
|  |      6 | to existing constructions (such as ``list of'') over the inductively defined
 | 
|  |      7 | set.  The examples below show how to use the additional expressiveness
 | 
|  |      8 | and how to reason from the resulting definitions.
 | 
| 10879 |      9 | 
 | 
|  |     10 | \subsection{Universal Quantifiers in Introduction Rules}
 | 
|  |     11 | \label{sec:gterm-datatype}
 | 
|  |     12 | 
 | 
| 11411 |     13 | \index{ground terms example|(}%
 | 
|  |     14 | \index{quantifiers!and inductive definitions|(}%
 | 
| 10879 |     15 | As a running example, this section develops the theory of \textbf{ground
 | 
|  |     16 | terms}: terms constructed from constant and function 
 | 
|  |     17 | symbols but not variables. To simplify matters further, we regard a
 | 
|  |     18 | constant as a function applied to the null argument  list.  Let us declare a
 | 
|  |     19 | datatype \isa{gterm} for the type of ground  terms. It is a type constructor
 | 
|  |     20 | whose argument is a type of  function symbols. 
 | 
|  |     21 | \begin{isabelle}
 | 
|  |     22 | \isacommand{datatype}\ 'f\ gterm\ =\ Apply\ 'f\ "'f\ gterm\ list"
 | 
|  |     23 | \end{isabelle}
 | 
|  |     24 | To try it out, we declare a datatype of some integer operations: 
 | 
|  |     25 | integer constants, the unary minus operator and the addition 
 | 
|  |     26 | operator. 
 | 
|  |     27 | \begin{isabelle}
 | 
|  |     28 | \isacommand{datatype}\ integer_op\ =\ Number\ int\ |\ UnaryMinus\ |\ Plus
 | 
|  |     29 | \end{isabelle}
 | 
|  |     30 | Now the type \isa{integer\_op gterm} denotes the ground 
 | 
|  |     31 | terms built over those symbols.
 | 
|  |     32 | 
 | 
|  |     33 | The type constructor \texttt{gterm} can be generalized to a function 
 | 
|  |     34 | over sets.  It returns 
 | 
|  |     35 | the set of ground terms that can be formed over a set \isa{F} of function symbols. For
 | 
|  |     36 | example,  we could consider the set of ground terms formed from the finite 
 | 
| 10889 |     37 | set \isa{\isacharbraceleft Number 2, UnaryMinus,
 | 
|  |     38 | Plus\isacharbraceright}.
 | 
| 10879 |     39 | 
 | 
|  |     40 | This concept is inductive. If we have a list \isa{args} of ground terms 
 | 
|  |     41 | over~\isa{F} and a function symbol \isa{f} in \isa{F}, then we 
 | 
|  |     42 | can apply \isa{f} to  \isa{args} to obtain another ground term. 
 | 
|  |     43 | The only difficulty is that the argument list may be of any length. Hitherto, 
 | 
|  |     44 | each rule in an inductive definition referred to the inductively 
 | 
|  |     45 | defined set a fixed number of times, typically once or twice. 
 | 
|  |     46 | A universal quantifier in the premise of the introduction rule 
 | 
|  |     47 | expresses that every element of \isa{args} belongs
 | 
|  |     48 | to our inductively defined set: is a ground term 
 | 
| 11173 |     49 | over~\isa{F}.  The function \isa{set} denotes the set of elements in a given 
 | 
| 10879 |     50 | list. 
 | 
|  |     51 | \begin{isabelle}
 | 
|  |     52 | \isacommand{consts}\ gterms\ ::\ "'f\ set\ \isasymRightarrow \ 'f\ gterm\ set"\isanewline
 | 
|  |     53 | \isacommand{inductive}\ "gterms\ F"\isanewline
 | 
|  |     54 | \isakeyword{intros}\isanewline
 | 
|  |     55 | step[intro!]:\ "\isasymlbrakk \isasymforall t\ \isasymin \ set\ args.\ t\ \isasymin \ gterms\ F;\ \ f\ \isasymin \ F\isasymrbrakk \isanewline
 | 
|  |     56 | \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isasymLongrightarrow \ (Apply\ f\ args)\ \isasymin \ gterms\
 | 
|  |     57 | F"
 | 
|  |     58 | \end{isabelle}
 | 
|  |     59 | 
 | 
|  |     60 | To demonstrate a proof from this definition, let us 
 | 
|  |     61 | show that the function \isa{gterms}
 | 
| 11173 |     62 | is \textbf{monotone}.  We shall need this concept shortly.
 | 
| 10879 |     63 | \begin{isabelle}
 | 
|  |     64 | \isacommand{lemma}\ gterms_mono:\ "F\isasymsubseteq G\ \isasymLongrightarrow \ gterms\ F\ \isasymsubseteq \ gterms\ G"\isanewline
 | 
|  |     65 | \isacommand{apply}\ clarify\isanewline
 | 
|  |     66 | \isacommand{apply}\ (erule\ gterms.induct)\isanewline
 | 
|  |     67 | \isacommand{apply}\ blast\isanewline
 | 
|  |     68 | \isacommand{done}
 | 
|  |     69 | \end{isabelle}
 | 
|  |     70 | Intuitively, this theorem says that
 | 
|  |     71 | enlarging the set of function symbols enlarges the set of ground 
 | 
|  |     72 | terms. The proof is a trivial rule induction.
 | 
|  |     73 | First we use the \isa{clarify} method to assume the existence of an element of
 | 
| 11147 |     74 | \isa{gterms~F}.  (We could have used \isa{intro subsetI}.)  We then
 | 
| 10879 |     75 | apply rule induction. Here is the resulting subgoal: 
 | 
|  |     76 | \begin{isabelle}
 | 
|  |     77 | \ 1.\ \isasymAnd x\ args\ f.\isanewline
 | 
|  |     78 | \ \ \ \ \ \ \ \isasymlbrakk F\ \isasymsubseteq \ G;\ \isasymforall t\isasymin set\ args.\ t\ \isasymin \ gterms\ F\ \isasymand \ t\ \isasymin \ gterms\ G;\ f\ \isasymin \ F\isasymrbrakk \isanewline
 | 
|  |     79 | \ \ \ \ \ \ \ \isasymLongrightarrow \ Apply\ f\ args\ \isasymin \ gterms\ G%
 | 
|  |     80 | \end{isabelle}
 | 
|  |     81 | %
 | 
|  |     82 | The assumptions state that \isa{f} belongs 
 | 
|  |     83 | to~\isa{F}, which is included in~\isa{G}, and that every element of the list \isa{args} is
 | 
|  |     84 | a ground term over~\isa{G}.  The \isa{blast} method finds this chain of reasoning easily.  
 | 
|  |     85 | 
 | 
|  |     86 | \begin{warn}
 | 
|  |     87 | Why do we call this function \isa{gterms} instead 
 | 
|  |     88 | of {\isa{gterm}}?  A constant may have the same name as a type.  However,
 | 
|  |     89 | name  clashes could arise in the theorems that Isabelle generates. 
 | 
|  |     90 | Our choice of names keeps {\isa{gterms.induct}} separate from 
 | 
|  |     91 | {\isa{gterm.induct}}.
 | 
|  |     92 | \end{warn}
 | 
|  |     93 | 
 | 
| 11173 |     94 | Call a term \textbf{well-formed} if each symbol occurring in it is applied
 | 
|  |     95 | to the correct number of arguments.  (This number is called the symbol's
 | 
|  |     96 | \textbf{arity}.)  We can express well-formedness by
 | 
|  |     97 | generalizing the inductive definition of
 | 
|  |     98 | \isa{gterms}.
 | 
|  |     99 | Suppose we are given a function called \isa{arity}, specifying the arities
 | 
|  |    100 | of all symbols.  In the inductive step, we have a list \isa{args} of such
 | 
|  |    101 | terms and a function  symbol~\isa{f}. If the length of the list matches the
 | 
|  |    102 | function's arity  then applying \isa{f} to \isa{args} yields a well-formed
 | 
|  |    103 | term. 
 | 
| 10879 |    104 | \begin{isabelle}
 | 
|  |    105 | \isacommand{consts}\ well_formed_gterm\ ::\ "('f\ \isasymRightarrow \ nat)\ \isasymRightarrow \ 'f\ gterm\ set"\isanewline
 | 
|  |    106 | \isacommand{inductive}\ "well_formed_gterm\ arity"\isanewline
 | 
|  |    107 | \isakeyword{intros}\isanewline
 | 
|  |    108 | step[intro!]:\ "\isasymlbrakk \isasymforall t\ \isasymin \ set\ args.\ t\ \isasymin \ well_formed_gterm\ arity;\ \ \isanewline
 | 
|  |    109 | \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ length\ args\ =\ arity\ f\isasymrbrakk \isanewline
 | 
|  |    110 | \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isasymLongrightarrow \ (Apply\ f\ args)\ \isasymin \ well_formed_gterm\
 | 
|  |    111 | arity"
 | 
|  |    112 | \end{isabelle}
 | 
|  |    113 | %
 | 
|  |    114 | The inductive definition neatly captures the reasoning above.
 | 
| 11173 |    115 | The universal quantification over the
 | 
| 11411 |    116 | \isa{set} of arguments expresses that all of them are well-formed.%
 | 
|  |    117 | \index{quantifiers!and inductive definitions|)}
 | 
| 10879 |    118 | 
 | 
| 11173 |    119 | 
 | 
| 11216 |    120 | \subsection{Alternative Definition Using a Monotone Function}
 | 
| 10879 |    121 | 
 | 
| 11411 |    122 | \index{monotone functions!and inductive definitions|(}% 
 | 
|  |    123 | An inductive definition may refer to the
 | 
|  |    124 | inductively defined  set through an arbitrary monotone function.  To
 | 
|  |    125 | demonstrate this powerful feature, let us
 | 
| 10879 |    126 | change the  inductive definition above, replacing the
 | 
|  |    127 | quantifier by a use of the function \isa{lists}. This
 | 
|  |    128 | function, from the Isabelle theory of lists, is analogous to the
 | 
|  |    129 | function \isa{gterms} declared above: if \isa{A} is a set then
 | 
|  |    130 | {\isa{lists A}} is the set of lists whose elements belong to
 | 
|  |    131 | \isa{A}.  
 | 
|  |    132 | 
 | 
|  |    133 | In the inductive definition of well-formed terms, examine the one
 | 
|  |    134 | introduction rule.  The first premise states that \isa{args} belongs to
 | 
|  |    135 | the \isa{lists} of well-formed terms.  This formulation is more
 | 
|  |    136 | direct, if more obscure, than using a universal quantifier.
 | 
|  |    137 | \begin{isabelle}
 | 
|  |    138 | \isacommand{consts}\ well_formed_gterm'\ ::\ "('f\ \isasymRightarrow \ nat)\ \isasymRightarrow \ 'f\ gterm\ set"\isanewline
 | 
|  |    139 | \isacommand{inductive}\ "well_formed_gterm'\ arity"\isanewline
 | 
|  |    140 | \isakeyword{intros}\isanewline
 | 
|  |    141 | step[intro!]:\ "\isasymlbrakk args\ \isasymin \ lists\ (well_formed_gterm'\ arity);\ \ \isanewline
 | 
|  |    142 | \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ length\ args\ =\ arity\ f\isasymrbrakk \isanewline
 | 
|  |    143 | \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isasymLongrightarrow \ (Apply\ f\ args)\ \isasymin \ well_formed_gterm'\ arity"\isanewline
 | 
|  |    144 | \isakeyword{monos}\ lists_mono
 | 
|  |    145 | \end{isabelle}
 | 
|  |    146 | 
 | 
| 12333 |    147 | We cite the theorem \isa{lists_mono} to justify 
 | 
|  |    148 | using the function \isa{lists}.%
 | 
|  |    149 | \footnote{This particular theorem is installed by default already, but we
 | 
|  |    150 | include the \isakeyword{monos} declaration in order to illustrate its syntax.}
 | 
| 10879 |    151 | \begin{isabelle}
 | 
|  |    152 | A\ \isasymsubseteq\ B\ \isasymLongrightarrow \ lists\ A\ \isasymsubseteq
 | 
| 11421 |    153 | \ lists\ B\rulenamedx{lists_mono}
 | 
| 10879 |    154 | \end{isabelle}
 | 
|  |    155 | %
 | 
| 11173 |    156 | Why must the function be monotone?  An inductive definition describes
 | 
| 10879 |    157 | an iterative construction: each element of the set is constructed by a
 | 
|  |    158 | finite number of introduction rule applications.  For example, the
 | 
|  |    159 | elements of \isa{even} are constructed by finitely many applications of
 | 
|  |    160 | the rules 
 | 
|  |    161 | \begin{isabelle}
 | 
|  |    162 | 0\ \isasymin \ even\isanewline
 | 
|  |    163 | n\ \isasymin \ even\ \isasymLongrightarrow \ (Suc\ (Suc\ n))\ \isasymin
 | 
|  |    164 | \ even
 | 
|  |    165 | \end{isabelle}
 | 
|  |    166 | All references to a set in its
 | 
|  |    167 | inductive definition must be positive.  Applications of an
 | 
|  |    168 | introduction rule cannot invalidate previous applications, allowing the
 | 
|  |    169 | construction process to converge.
 | 
|  |    170 | The following pair of rules do not constitute an inductive definition:
 | 
|  |    171 | \begin{isabelle}
 | 
|  |    172 | 0\ \isasymin \ even\isanewline
 | 
|  |    173 | n\ \isasymnotin \ even\ \isasymLongrightarrow \ (Suc\ n)\ \isasymin
 | 
|  |    174 | \ even
 | 
|  |    175 | \end{isabelle}
 | 
|  |    176 | %
 | 
|  |    177 | Showing that 4 is even using these rules requires showing that 3 is not
 | 
|  |    178 | even.  It is far from trivial to show that this set of rules
 | 
|  |    179 | characterizes the even numbers.  
 | 
|  |    180 | 
 | 
|  |    181 | Even with its use of the function \isa{lists}, the premise of our
 | 
|  |    182 | introduction rule is positive:
 | 
|  |    183 | \begin{isabelle}
 | 
|  |    184 | args\ \isasymin \ lists\ (well_formed_gterm'\ arity)
 | 
|  |    185 | \end{isabelle}
 | 
|  |    186 | To apply the rule we construct a list \isa{args} of previously
 | 
|  |    187 | constructed well-formed terms.  We obtain a
 | 
| 11173 |    188 | new term, \isa{Apply\ f\ args}.  Because \isa{lists} is monotone,
 | 
| 10879 |    189 | applications of the rule remain valid as new terms are constructed.
 | 
|  |    190 | Further lists of well-formed
 | 
| 11411 |    191 | terms become available and none are taken away.%
 | 
|  |    192 | \index{monotone functions!and inductive definitions|)} 
 | 
| 10879 |    193 | 
 | 
|  |    194 | 
 | 
|  |    195 | \subsection{A Proof of Equivalence}
 | 
|  |    196 | 
 | 
|  |    197 | We naturally hope that these two inductive definitions of ``well-formed'' 
 | 
|  |    198 | coincide.  The equality can be proved by separate inclusions in 
 | 
|  |    199 | each direction.  Each is a trivial rule induction. 
 | 
|  |    200 | \begin{isabelle}
 | 
|  |    201 | \isacommand{lemma}\ "well_formed_gterm\ arity\ \isasymsubseteq \ well_formed_gterm'\ arity"\isanewline
 | 
|  |    202 | \isacommand{apply}\ clarify\isanewline
 | 
|  |    203 | \isacommand{apply}\ (erule\ well_formed_gterm.induct)\isanewline
 | 
|  |    204 | \isacommand{apply}\ auto\isanewline
 | 
|  |    205 | \isacommand{done}
 | 
|  |    206 | \end{isabelle}
 | 
|  |    207 | 
 | 
|  |    208 | The \isa{clarify} method gives
 | 
|  |    209 | us an element of \isa{well_formed_gterm\ arity} on which to perform 
 | 
|  |    210 | induction.  The resulting subgoal can be proved automatically:
 | 
|  |    211 | \begin{isabelle}
 | 
|  |    212 | {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ args\ f{\isachardot}\isanewline
 | 
|  |    213 | \ \ \ \ \ \ {\isasymlbrakk}{\isasymforall}t{\isasymin}set\ args{\isachardot}\isanewline
 | 
| 11421 |    214 | \ \ \ \ \ \ \ \ \ \ t\ {\isasymin}\ well\_formed\_gterm\ arity\ {\isasymand}\ t\ {\isasymin}\ well\_formed\_gterm{\isacharprime}\ arity{\isacharsemicolon}\isanewline
 | 
| 10879 |    215 | \ \ \ \ \ \ \ length\ args\ {\isacharequal}\ arity\ f{\isasymrbrakk}\isanewline
 | 
| 11421 |    216 | \ \ \ \ \ \ {\isasymLongrightarrow}\ Apply\ f\ args\ {\isasymin}\ well\_formed\_gterm{\isacharprime}\ arity%
 | 
| 10879 |    217 | \end{isabelle}
 | 
|  |    218 | %
 | 
|  |    219 | This proof resembles the one given in
 | 
|  |    220 | {\S}\ref{sec:gterm-datatype} above, especially in the form of the
 | 
|  |    221 | induction hypothesis.  Next, we consider the opposite inclusion:
 | 
|  |    222 | \begin{isabelle}
 | 
|  |    223 | \isacommand{lemma}\ "well_formed_gterm'\ arity\ \isasymsubseteq \ well_formed_gterm\ arity"\isanewline
 | 
|  |    224 | \isacommand{apply}\ clarify\isanewline
 | 
|  |    225 | \isacommand{apply}\ (erule\ well_formed_gterm'.induct)\isanewline
 | 
|  |    226 | \isacommand{apply}\ auto\isanewline
 | 
|  |    227 | \isacommand{done}
 | 
|  |    228 | \end{isabelle}
 | 
|  |    229 | %
 | 
|  |    230 | The proof script is identical, but the subgoal after applying induction may
 | 
|  |    231 | be surprising:
 | 
|  |    232 | \begin{isabelle}
 | 
|  |    233 | 1.\ \isasymAnd x\ args\ f.\isanewline
 | 
|  |    234 | \ \ \ \ \ \ \isasymlbrakk args\ \isasymin \ lists\ (well_formed_gterm'\ arity\ \isasyminter\isanewline
 | 
|  |    235 | \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacharbraceleft u.\ u\ \isasymin \ well_formed_gterm\ arity\isacharbraceright );\isanewline
 | 
|  |    236 | \ \ \ \ \ \ \ length\ args\ =\ arity\ f\isasymrbrakk \isanewline
 | 
|  |    237 | \ \ \ \ \ \ \isasymLongrightarrow \ Apply\ f\ args\ \isasymin \ well_formed_gterm\ arity%
 | 
|  |    238 | \end{isabelle}
 | 
|  |    239 | The induction hypothesis contains an application of \isa{lists}.  Using a
 | 
| 11173 |    240 | monotone function in the inductive definition always has this effect.  The
 | 
|  |    241 | subgoal may look uninviting, but fortunately 
 | 
|  |    242 | \isa{lists} distributes over intersection:
 | 
| 10879 |    243 | \begin{isabelle}
 | 
|  |    244 | lists\ (A\ \isasyminter \ B)\ =\ lists\ A\ \isasyminter \ lists\ B
 | 
|  |    245 | \rulename{lists_Int_eq}
 | 
|  |    246 | \end{isabelle}
 | 
|  |    247 | %
 | 
|  |    248 | Thanks to this default simplification rule, the induction hypothesis 
 | 
|  |    249 | is quickly replaced by its two parts:
 | 
|  |    250 | \begin{isabelle}
 | 
|  |    251 | \ \ \ \ \ \ \ args\ \isasymin \ lists\ (well_formed_gterm'\ arity)\isanewline
 | 
|  |    252 | \ \ \ \ \ \ \ args\ \isasymin \ lists\ (well_formed_gterm\ arity)
 | 
|  |    253 | \end{isabelle}
 | 
|  |    254 | Invoking the rule \isa{well_formed_gterm.step} completes the proof.  The
 | 
|  |    255 | call to
 | 
|  |    256 | \isa{auto} does all this work.
 | 
|  |    257 | 
 | 
| 11411 |    258 | This example is typical of how monotone functions
 | 
|  |    259 | \index{monotone functions} can be used.  In particular, many of them
 | 
|  |    260 | distribute over intersection.  Monotonicity implies one direction of
 | 
|  |    261 | this set equality; we have this theorem:
 | 
| 10879 |    262 | \begin{isabelle}
 | 
|  |    263 | mono\ f\ \isasymLongrightarrow \ f\ (A\ \isasyminter \ B)\ \isasymsubseteq \
 | 
|  |    264 | f\ A\ \isasyminter \ f\ B%
 | 
|  |    265 | \rulename{mono_Int}
 | 
|  |    266 | \end{isabelle}
 | 
|  |    267 | 
 | 
|  |    268 | 
 | 
| 11173 |    269 | \subsection{Another Example of Rule Inversion}
 | 
|  |    270 | 
 | 
| 11411 |    271 | \index{rule inversion|(}%
 | 
| 11173 |    272 | Does \isa{gterms} distribute over intersection?  We have proved that this
 | 
|  |    273 | function is monotone, so \isa{mono_Int} gives one of the inclusions.  The
 | 
|  |    274 | opposite inclusion asserts that if \isa{t} is a ground term over both of the
 | 
|  |    275 | sets
 | 
|  |    276 | \isa{F} and~\isa{G} then it is also a ground term over their intersection,
 | 
|  |    277 | \isa{F\isasyminter G}.
 | 
| 11261 |    278 | \begin{isabelle}
 | 
|  |    279 | \isacommand{lemma}\ gterms_IntI:\isanewline
 | 
|  |    280 | \ \ \ \ \ "t\ \isasymin \ gterms\ F\ \isasymLongrightarrow \ t\ \isasymin \ gterms\ G\ \isasymlongrightarrow \ t\ \isasymin \ gterms\ (F\isasyminter
 | 
|  |    281 | G)"
 | 
|  |    282 | \end{isabelle}
 | 
| 11173 |    283 | Attempting this proof, we get the assumption 
 | 
| 11261 |    284 | \isa{Apply\ f\ args\ \isasymin\ gterms\ G}, which cannot be broken down. 
 | 
| 11173 |    285 | It looks like a job for rule inversion:
 | 
|  |    286 | \begin{isabelle}
 | 
| 11411 |    287 | \commdx{inductive\protect\_cases}\ gterm_Apply_elim\ [elim!]:\ "Apply\ f\ args\
 | 
| 11173 |    288 | \isasymin\ gterms\ F"
 | 
|  |    289 | \end{isabelle}
 | 
|  |    290 | %
 | 
|  |    291 | Here is the result.
 | 
|  |    292 | \begin{isabelle}
 | 
|  |    293 | \isasymlbrakk Apply\ f\ args\ \isasymin \ gterms\ F;\isanewline
 | 
|  |    294 | \ \isasymlbrakk
 | 
|  |    295 | \isasymforall t\isasymin set\ args.\ t\ \isasymin \ gterms\ F;\ f\ \isasymin
 | 
|  |    296 | \ F\isasymrbrakk \ \isasymLongrightarrow \ P\isasymrbrakk\isanewline
 | 
|  |    297 | \isasymLongrightarrow \ P%
 | 
|  |    298 | \rulename{gterm_Apply_elim}
 | 
|  |    299 | \end{isabelle}
 | 
|  |    300 | This rule replaces an assumption about \isa{Apply\ f\ args} by 
 | 
|  |    301 | assumptions about \isa{f} and~\isa{args}.  
 | 
|  |    302 | No cases are discarded (there was only one to begin
 | 
|  |    303 | with) but the rule applies specifically to the pattern \isa{Apply\ f\ args}.
 | 
|  |    304 | It can be applied repeatedly as an elimination rule without looping, so we
 | 
|  |    305 | have given the
 | 
|  |    306 | \isa{elim!}\ attribute. 
 | 
| 10879 |    307 | 
 | 
| 11173 |    308 | Now we can prove the other half of that distributive law.
 | 
|  |    309 | \begin{isabelle}
 | 
|  |    310 | \isacommand{lemma}\ gterms_IntI\ [rule_format,\ intro!]:\isanewline
 | 
|  |    311 | \ \ \ \ \ "t\ \isasymin \ gterms\ F\ \isasymLongrightarrow \ t\ \isasymin \ gterms\ G\ \isasymlongrightarrow \ t\ \isasymin \ gterms\ (F\isasyminter G)"\isanewline
 | 
|  |    312 | \isacommand{apply}\ (erule\ gterms.induct)\isanewline
 | 
|  |    313 | \isacommand{apply}\ blast\isanewline
 | 
|  |    314 | \isacommand{done}
 | 
|  |    315 | \end{isabelle}
 | 
|  |    316 | %
 | 
|  |    317 | The proof begins with rule induction over the definition of
 | 
|  |    318 | \isa{gterms}, which leaves a single subgoal:  
 | 
|  |    319 | \begin{isabelle}
 | 
|  |    320 | 1.\ \isasymAnd args\ f.\isanewline
 | 
|  |    321 | \ \ \ \ \ \ \isasymlbrakk \isasymforall t\isasymin set\ args.\ t\ \isasymin \ gterms\ F\ \isasymand\isanewline
 | 
|  |    322 | \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ (t\ \isasymin \ gterms\ G\ \isasymlongrightarrow \ t\ \isasymin \ gterms\ (F\ \isasyminter \ G));\isanewline
 | 
|  |    323 | \ \ \ \ \ \ \ f\ \isasymin \ F\isasymrbrakk \isanewline
 | 
|  |    324 | \ \ \ \ \ \ \isasymLongrightarrow \ Apply\ f\ args\ \isasymin \ gterms\ G\ \isasymlongrightarrow \ Apply\ f\ args\ \isasymin \ gterms\ (F\ \isasyminter \ G)
 | 
|  |    325 | \end{isabelle}
 | 
|  |    326 | %
 | 
|  |    327 | To prove this, we assume \isa{Apply\ f\ args\ \isasymin \
 | 
|  |    328 | gterms\ G}.  Rule inversion, in the form of \isa{gterm_Apply_elim}, infers
 | 
|  |    329 | that every element of \isa{args} belongs to 
 | 
|  |    330 | \isa{gterms~G}; hence (by the induction hypothesis) it belongs
 | 
|  |    331 | to \isa{gterms\ (F\ \isasyminter \ G)}.  Rule inversion also yields
 | 
|  |    332 | \isa{f\ \isasymin\ G} and hence \isa{f\ \isasymin \ F\ \isasyminter \ G}. 
 | 
|  |    333 | All of this reasoning is done by \isa{blast}.
 | 
|  |    334 | 
 | 
|  |    335 | \smallskip
 | 
|  |    336 | Our distributive law is a trivial consequence of previously-proved results:
 | 
|  |    337 | \begin{isabelle}
 | 
|  |    338 | \isacommand{theorem}\ gterms_Int_eq\ [simp]:\isanewline
 | 
|  |    339 | \ \ \ \ \ "gterms\ (F\isasyminter G)\ =\ gterms\ F\ \isasyminter \ gterms\ G"\isanewline
 | 
|  |    340 | \isacommand{by}\ (blast\ intro!:\ mono_Int\ monoI\ gterms_mono)
 | 
|  |    341 | \end{isabelle}
 | 
| 11411 |    342 | \index{rule inversion|)}%
 | 
|  |    343 | \index{ground terms example|)}
 | 
| 10879 |    344 | 
 | 
|  |    345 | 
 | 
|  |    346 | \begin{exercise}
 | 
|  |    347 | A function mapping function symbols to their 
 | 
|  |    348 | types is called a \textbf{signature}.  Given a type 
 | 
|  |    349 | ranging over type symbols, we can represent a function's type by a
 | 
|  |    350 | list of argument types paired with the result type. 
 | 
|  |    351 | Complete this inductive definition:
 | 
|  |    352 | \begin{isabelle}
 | 
|  |    353 | \isacommand{consts}\ well_typed_gterm::\ "('f\ \isasymRightarrow \ 't\ list\ *\ 't)\ \isasymRightarrow \ ('f\ gterm\ *\ 't)set"\isanewline
 | 
|  |    354 | \isacommand{inductive}\ "well_typed_gterm\ sig"\isanewline
 | 
|  |    355 | \end{isabelle}
 | 
|  |    356 | \end{exercise}
 |