author | berghofe |
Wed, 11 Jul 2007 10:53:39 +0200 | |
changeset 23733 | 3f8ad7418e55 |
parent 16417 | 9bc16273c2d4 |
child 23842 | 9d87177f1f89 |
permissions | -rw-r--r-- |
10370 | 1 |
(* ID: $Id$ *) |
16417 | 2 |
theory Advanced imports Even begin |
10426 | 3 |
|
10468 | 4 |
|
5 |
datatype 'f gterm = Apply 'f "'f gterm list" |
|
6 |
||
7 |
datatype integer_op = Number int | UnaryMinus | Plus; |
|
8 |
||
23733 | 9 |
inductive_set |
10 |
gterms :: "'f set \<Rightarrow> 'f gterm set" |
|
11 |
for F :: "'f set" |
|
12 |
where |
|
10468 | 13 |
step[intro!]: "\<lbrakk>\<forall>t \<in> set args. t \<in> gterms F; f \<in> F\<rbrakk> |
14 |
\<Longrightarrow> (Apply f args) \<in> gterms F" |
|
15 |
||
16 |
lemma gterms_mono: "F\<subseteq>G \<Longrightarrow> gterms F \<subseteq> gterms G" |
|
17 |
apply clarify |
|
18 |
apply (erule gterms.induct) |
|
10882 | 19 |
txt{* |
20 |
@{subgoals[display,indent=0,margin=65]} |
|
21 |
*}; |
|
10468 | 22 |
apply blast |
23 |
done |
|
24 |
||
10882 | 25 |
|
10468 | 26 |
text{* |
27 |
@{thm[display] even.cases[no_vars]} |
|
28 |
\rulename{even.cases} |
|
29 |
||
30 |
Just as a demo I include |
|
31 |
the two forms that Markus has made available. First the one for binding the |
|
32 |
result to a name |
|
33 |
||
34 |
*} |
|
35 |
||
36 |
inductive_cases Suc_Suc_cases [elim!]: |
|
10426 | 37 |
"Suc(Suc n) \<in> even" |
38 |
||
10468 | 39 |
thm Suc_Suc_cases; |
40 |
||
41 |
text{* |
|
42 |
@{thm[display] Suc_Suc_cases[no_vars]} |
|
43 |
\rulename{Suc_Suc_cases} |
|
44 |
||
45 |
and now the one for local usage: |
|
46 |
*} |
|
47 |
||
48 |
lemma "Suc(Suc n) \<in> even \<Longrightarrow> P n"; |
|
49 |
apply (ind_cases "Suc(Suc n) \<in> even"); |
|
50 |
oops |
|
51 |
||
52 |
inductive_cases gterm_Apply_elim [elim!]: "Apply f args \<in> gterms F" |
|
53 |
||
54 |
text{*this is what we get: |
|
55 |
||
56 |
@{thm[display] gterm_Apply_elim[no_vars]} |
|
57 |
\rulename{gterm_Apply_elim} |
|
58 |
*} |
|
10426 | 59 |
|
11173 | 60 |
lemma gterms_IntI [rule_format, intro!]: |
10468 | 61 |
"t \<in> gterms F \<Longrightarrow> t \<in> gterms G \<longrightarrow> t \<in> gterms (F\<inter>G)" |
62 |
apply (erule gterms.induct) |
|
10882 | 63 |
txt{* |
64 |
@{subgoals[display,indent=0,margin=65]} |
|
65 |
*}; |
|
10468 | 66 |
apply blast |
67 |
done |
|
68 |
||
69 |
||
70 |
text{* |
|
71 |
@{thm[display] mono_Int[no_vars]} |
|
72 |
\rulename{mono_Int} |
|
73 |
*} |
|
10426 | 74 |
|
10468 | 75 |
lemma gterms_Int_eq [simp]: |
76 |
"gterms (F\<inter>G) = gterms F \<inter> gterms G" |
|
11173 | 77 |
by (blast intro!: mono_Int monoI gterms_mono) |
10468 | 78 |
|
79 |
||
12156
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
paulson
parents:
11711
diff
changeset
|
80 |
text{*the following declaration isn't actually used*} |
10468 | 81 |
consts integer_arity :: "integer_op \<Rightarrow> nat" |
82 |
primrec |
|
11711 | 83 |
"integer_arity (Number n) = 0" |
84 |
"integer_arity UnaryMinus = 1" |
|
85 |
"integer_arity Plus = 2" |
|
10468 | 86 |
|
23733 | 87 |
inductive_set |
88 |
well_formed_gterm :: "('f \<Rightarrow> nat) \<Rightarrow> 'f gterm set" |
|
89 |
for arity :: "'f \<Rightarrow> nat" |
|
90 |
where |
|
10468 | 91 |
step[intro!]: "\<lbrakk>\<forall>t \<in> set args. t \<in> well_formed_gterm arity; |
92 |
length args = arity f\<rbrakk> |
|
93 |
\<Longrightarrow> (Apply f args) \<in> well_formed_gterm arity" |
|
94 |
||
10426 | 95 |
|
23733 | 96 |
inductive_set |
97 |
well_formed_gterm' :: "('f \<Rightarrow> nat) \<Rightarrow> 'f gterm set" |
|
98 |
for arity :: "'f \<Rightarrow> nat" |
|
99 |
where |
|
10468 | 100 |
step[intro!]: "\<lbrakk>args \<in> lists (well_formed_gterm' arity); |
101 |
length args = arity f\<rbrakk> |
|
102 |
\<Longrightarrow> (Apply f args) \<in> well_formed_gterm' arity" |
|
103 |
monos lists_mono |
|
104 |
||
105 |
lemma "well_formed_gterm arity \<subseteq> well_formed_gterm' arity" |
|
106 |
apply clarify |
|
10882 | 107 |
txt{* |
108 |
The situation after clarify |
|
109 |
@{subgoals[display,indent=0,margin=65]} |
|
110 |
*}; |
|
10468 | 111 |
apply (erule well_formed_gterm.induct) |
10882 | 112 |
txt{* |
113 |
note the induction hypothesis! |
|
114 |
@{subgoals[display,indent=0,margin=65]} |
|
115 |
*}; |
|
10468 | 116 |
apply auto |
117 |
done |
|
118 |
||
10426 | 119 |
|
10370 | 120 |
|
10468 | 121 |
lemma "well_formed_gterm' arity \<subseteq> well_formed_gterm arity" |
122 |
apply clarify |
|
10882 | 123 |
txt{* |
124 |
The situation after clarify |
|
125 |
@{subgoals[display,indent=0,margin=65]} |
|
126 |
*}; |
|
10468 | 127 |
apply (erule well_formed_gterm'.induct) |
10882 | 128 |
txt{* |
129 |
note the induction hypothesis! |
|
130 |
@{subgoals[display,indent=0,margin=65]} |
|
131 |
*}; |
|
10468 | 132 |
apply auto |
133 |
done |
|
10370 | 134 |
|
135 |
||
10468 | 136 |
text{* |
137 |
@{thm[display] lists_Int_eq[no_vars]} |
|
138 |
*} |
|
139 |
||
140 |
text{* the rest isn't used: too complicated. OK for an exercise though.*} |
|
141 |
||
23733 | 142 |
inductive_set |
143 |
integer_signature :: "(integer_op * (unit list * unit)) set" |
|
144 |
where |
|
145 |
Number: "(Number n, ([], ())) \<in> integer_signature" |
|
146 |
| UnaryMinus: "(UnaryMinus, ([()], ())) \<in> integer_signature" |
|
147 |
| Plus: "(Plus, ([(),()], ())) \<in> integer_signature" |
|
10468 | 148 |
|
10370 | 149 |
|
23733 | 150 |
inductive_set |
151 |
well_typed_gterm :: "('f \<Rightarrow> 't list * 't) \<Rightarrow> ('f gterm * 't)set" |
|
152 |
for sig :: "'f \<Rightarrow> 't list * 't" |
|
153 |
where |
|
10468 | 154 |
step[intro!]: |
155 |
"\<lbrakk>\<forall>pair \<in> set args. pair \<in> well_typed_gterm sig; |
|
156 |
sig f = (map snd args, rtype)\<rbrakk> |
|
157 |
\<Longrightarrow> (Apply f (map fst args), rtype) |
|
158 |
\<in> well_typed_gterm sig" |
|
159 |
||
23733 | 160 |
inductive_set |
161 |
well_typed_gterm' :: "('f \<Rightarrow> 't list * 't) \<Rightarrow> ('f gterm * 't)set" |
|
162 |
for sig :: "'f \<Rightarrow> 't list * 't" |
|
163 |
where |
|
10468 | 164 |
step[intro!]: |
165 |
"\<lbrakk>args \<in> lists(well_typed_gterm' sig); |
|
166 |
sig f = (map snd args, rtype)\<rbrakk> |
|
167 |
\<Longrightarrow> (Apply f (map fst args), rtype) |
|
168 |
\<in> well_typed_gterm' sig" |
|
10370 | 169 |
monos lists_mono |
170 |
||
171 |
||
10468 | 172 |
lemma "well_typed_gterm sig \<subseteq> well_typed_gterm' sig" |
10370 | 173 |
apply clarify |
10468 | 174 |
apply (erule well_typed_gterm.induct) |
10370 | 175 |
apply auto |
176 |
done |
|
177 |
||
10468 | 178 |
lemma "well_typed_gterm' sig \<subseteq> well_typed_gterm sig" |
10370 | 179 |
apply clarify |
10468 | 180 |
apply (erule well_typed_gterm'.induct) |
10370 | 181 |
apply auto |
182 |
done |
|
183 |
||
10468 | 184 |
|
10370 | 185 |
end |
186 |