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