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