91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 1
theory Ind_Code
441
+ − 2
imports Ind_General_Scheme "../First_Steps"
91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 3
begin
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 4
565
+ − 5
section \<open>The Gory Details\label{sec:code}\<close>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 6
565
+ − 7
text \<open>
237
+ − 8
As mentioned before the code falls roughly into three parts: the code that deals
+ − 9
with the definitions, with the induction principles and with the introduction
+ − 10
rules. In addition there are some administrative functions that string everything
+ − 11
together.
565
+ − 12
\<close>
208
+ − 13
565
+ − 14
subsection \<open>Definitions\<close>
208
+ − 15
565
+ − 16
text \<open>
237
+ − 17
We first have to produce for each predicate the user specifies an appropriate
+ − 18
definition, whose general form is
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 19
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 20
@{text [display] "pred \<equiv> \<lambda>zs. \<forall>preds. orules \<longrightarrow> pred zs"}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 21
208
+ − 22
and then ``register'' the definition inside a local theory.
237
+ − 23
To do the latter, we use the following wrapper for the function
394
+ − 24
@{ML_ind define in Local_Theory}. The wrapper takes a predicate name, a syntax
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 25
annotation and a term representing the right-hand side of the definition.
565
+ − 26
\<close>
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 27
565
+ − 28
ML %linenosgray\<open>fun make_defn ((predname, mx), trm) lthy =
164
+ − 29
let
562
+ − 30
val arg = ((predname, mx), (Binding.empty_atts, trm))
401
+ − 31
val ((_, (_ , thm)), lthy') = Local_Theory.define arg lthy
164
+ − 32
in
176
+ − 33
(thm, lthy')
565
+ − 34
end\<close>
164
+ − 35
565
+ − 36
text \<open>
237
+ − 37
It returns the definition (as a theorem) and the local theory in which the
562
+ − 38
definition has been made. We use @{ML_ind empty_atts in Binding} in Line 3,
401
+ − 39
since the definitions for our inductive predicates are not meant to be seen
+ − 40
by the user and therefore do not need to have any theorem attributes.
+ − 41
208
+ − 42
The next two functions construct the right-hand sides of the definitions,
237
+ − 43
which are terms whose general form is:
179
+ − 44
190
+ − 45
@{text [display] "\<lambda>zs. \<forall>preds. orules \<longrightarrow> pred zs"}
179
+ − 46
565
+ − 47
When constructing these terms, the variables \<open>zs\<close> need to be chosen so
+ − 48
that they do not occur in the \<open>orules\<close> and also be distinct from the
+ − 49
\<open>preds\<close>.
208
+ − 50
183
+ − 51
565
+ − 52
The first function, named \<open>defn_aux\<close>, constructs the term for one
+ − 53
particular predicate (the argument \<open>pred\<close> in the code below). The
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 54
number of arguments of this predicate is determined by the number of
565
+ − 55
argument types given in \<open>arg_tys\<close>. The other arguments of the
+ − 56
function are the \<open>orules\<close> and all the \<open>preds\<close>.
+ − 57
\<close>
164
+ − 58
565
+ − 59
ML %linenosgray\<open>fun defn_aux lthy orules preds (pred, arg_tys) =
164
+ − 60
let
+ − 61
fun mk_all x P = HOLogic.all_const (fastype_of x) $ lambda x P
+ − 62
+ − 63
val fresh_args =
176
+ − 64
arg_tys
164
+ − 65
|> map (pair "z")
179
+ − 66
|> Variable.variant_frees lthy (preds @ orules)
164
+ − 67
|> map Free
+ − 68
in
+ − 69
list_comb (pred, fresh_args)
+ − 70
|> fold_rev (curry HOLogic.mk_imp) orules
+ − 71
|> fold_rev mk_all preds
+ − 72
|> fold_rev lambda fresh_args
565
+ − 73
end\<close>
164
+ − 74
565
+ − 75
text \<open>
+ − 76
The function \<open>mk_all\<close> in Line 3 is just a helper function for constructing
+ − 77
universal quantifications. The code in Lines 5 to 9 produces the fresh \<open>zs\<close>. For this it pairs every argument type with the string
184
+ − 78
@{text [quotes] "z"} (Line 7); then generates variants for all these strings
565
+ − 79
so that they are unique w.r.t.~to the predicates and \<open>orules\<close> (Line 8);
184
+ − 80
in Line 9 it generates the corresponding variable terms for the unique
+ − 81
strings.
164
+ − 82
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 83
The unique variables are applied to the predicate in Line 11 using the
565
+ − 84
function @{ML list_comb}; then the \<open>orules\<close> are prefixed (Line 12); in
183
+ − 85
Line 13 we quantify over all predicates; and in line 14 we just abstract
565
+ − 86
over all the \<open>zs\<close>, i.e., the fresh arguments of the
190
+ − 87
predicate. A testcase for this function is
565
+ − 88
\<close>
164
+ − 89
565
+ − 90
local_setup %gray \<open>fn lthy =>
176
+ − 91
let
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 92
val def = defn_aux lthy eo_orules eo_preds (e_pred, e_arg_tys)
173
+ − 93
in
441
+ − 94
pwriteln (pretty_term lthy def); lthy
565
+ − 95
end\<close>
173
+ − 96
565
+ − 97
text \<open>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 98
where we use the shorthands defined in Figure~\ref{fig:shorthands}.
565
+ − 99
The testcase calls @{ML defn_aux} for the predicate \<open>even\<close> and prints
208
+ − 100
out the generated definition. So we obtain as printout
179
+ − 101
+ − 102
@{text [display]
+ − 103
"\<lambda>z. \<forall>even odd. (even 0) \<longrightarrow> (\<forall>n. odd n \<longrightarrow> even (Suc n))
+ − 104
\<longrightarrow> (\<forall>n. even n \<longrightarrow> odd (Suc n)) \<longrightarrow> even z"}
+ − 105
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 106
If we try out the function with the rules for freshness
565
+ − 107
\<close>
210
+ − 108
565
+ − 109
local_setup %gray \<open>fn lthy =>
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 110
let
358
+ − 111
val arg = (fresh_pred, fresh_arg_tys)
+ − 112
val def = defn_aux lthy fresh_orules [fresh_pred] arg
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 113
in
441
+ − 114
pwriteln (pretty_term lthy def); lthy
565
+ − 115
end\<close>
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 116
210
+ − 117
565
+ − 118
text \<open>
210
+ − 119
we obtain
+ − 120
+ − 121
@{term [display]
+ − 122
"\<lambda>z za. \<forall>fresh. (\<forall>a b. \<not> a = b \<longrightarrow> fresh a (Var b)) \<longrightarrow>
+ − 123
(\<forall>a s t. fresh a t \<longrightarrow> fresh a s \<longrightarrow> fresh a (App t s)) \<longrightarrow>
+ − 124
(\<forall>a t. fresh a (Lam a t)) \<longrightarrow>
+ − 125
(\<forall>a b t. \<not> a = b \<longrightarrow> fresh a t \<longrightarrow> fresh a (Lam b t)) \<longrightarrow> fresh z za"}
+ − 126
+ − 127
565
+ − 128
The second function, named \<open>defns\<close>, has to iterate the function
+ − 129
@{ML defn_aux} over all predicates. The argument \<open>preds\<close> is again
+ − 130
the list of predicates as @{ML_type term}s; the argument \<open>prednames\<close> is the list of binding names of the predicates; \<open>mxs\<close>
237
+ − 131
are the list of syntax, or mixfix, annotations for the predicates;
565
+ − 132
\<open>arg_tyss\<close> is the list of argument-type-lists.
+ − 133
\<close>
164
+ − 134
565
+ − 135
ML %linenosgray\<open>fun defns rules preds prednames mxs arg_typss lthy =
164
+ − 136
let
562
+ − 137
val orules = map (Object_Logic.atomize_term lthy) rules
210
+ − 138
val defs = map (defn_aux lthy orules preds) (preds ~~ arg_typss)
164
+ − 139
in
237
+ − 140
fold_map make_defn (prednames ~~ mxs ~~ defs) lthy
565
+ − 141
end\<close>
164
+ − 142
565
+ − 143
text \<open>
184
+ − 144
The user will state the introduction rules using meta-implications and
575
+ − 145
meta-quanti\-fications. In Line 3, we transform these introduction rules
237
+ − 146
into the object logic (since definitions cannot be stated with
575
+ − 147
meta-connectives). We use the function
+ − 148
@{ML_ind atomize_term in Object_Logic} to make the transformation (Line 3). The call
+ − 149
to @{ML defn_aux} in Line 4 produces all right-hand sides of the
+ − 150
definitions. The actual definitions are then made in Line 6. The result of
237
+ − 151
the function is a list of theorems and a local theory (the theorems are
+ − 152
registered with the local theory). A testcase for this function is
565
+ − 153
\<close>
164
+ − 154
565
+ − 155
local_setup %gray \<open>fn lthy =>
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 156
let
208
+ − 157
val (defs, lthy') =
237
+ − 158
defns eo_rules eo_preds eo_prednames eo_mxs eo_arg_tyss lthy
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 159
in
440
+ − 160
pwriteln (pretty_thms_no_vars lthy' defs); lthy
565
+ − 161
end\<close>
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 162
565
+ − 163
text \<open>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 164
where we feed into the function all parameters corresponding to
565
+ − 165
the \<open>even\<close>/\<open>odd\<close> example. The definitions we obtain
184
+ − 166
are:
179
+ − 167
210
+ − 168
@{text [display, break]
+ − 169
"even \<equiv> \<lambda>z. \<forall>even odd. (even 0) \<longrightarrow> (\<forall>n. odd n \<longrightarrow> even (Suc n))
+ − 170
\<longrightarrow> (\<forall>n. even n \<longrightarrow> odd (Suc n)) \<longrightarrow> even z,
+ − 171
odd \<equiv> \<lambda>z. \<forall>even odd. (even 0) \<longrightarrow> (\<forall>n. odd n \<longrightarrow> even (Suc n))
+ − 172
\<longrightarrow> (\<forall>n. even n \<longrightarrow> odd (Suc n)) \<longrightarrow> odd z"}
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 173
565
+ − 174
Note that in the testcase we return the local theory \<open>lthy\<close>
+ − 175
(not the modified \<open>lthy'\<close>). As a result the test case has no effect
210
+ − 176
on the ambient theory. The reason is that if we introduce the
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 177
definition again, we pollute the name space with two versions of
565
+ − 178
\<open>even\<close> and \<open>odd\<close>. We want to avoid this here.
164
+ − 179
210
+ − 180
This completes the code for introducing the definitions. Next we deal with
208
+ − 181
the induction principles.
565
+ − 182
\<close>
208
+ − 183
565
+ − 184
subsection \<open>Induction Principles\<close>
208
+ − 185
565
+ − 186
text \<open>
215
+ − 187
Recall that the manual proof for the induction principle
565
+ − 188
of \<open>even\<close> was:
+ − 189
\<close>
176
+ − 190
210
+ − 191
lemma manual_ind_prin_even:
209
+ − 192
assumes prem: "even z"
+ − 193
shows "P 0 \<Longrightarrow> (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P z"
176
+ − 194
apply(atomize (full))
209
+ − 195
apply(cut_tac prem)
176
+ − 196
apply(unfold even_def)
+ − 197
apply(drule spec[where x=P])
+ − 198
apply(drule spec[where x=Q])
+ − 199
apply(assumption)
+ − 200
done
+ − 201
565
+ − 202
text \<open>
190
+ − 203
The code for automating such induction principles has to accomplish two tasks:
184
+ − 204
constructing the induction principles from the given introduction
190
+ − 205
rules and then automatically generating proofs for them using a tactic.
183
+ − 206
+ − 207
The tactic will use the following helper function for instantiating universal
184
+ − 208
quantifiers.
565
+ − 209
\<close>
176
+ − 210
565
+ − 211
ML %grayML\<open>fun inst_spec ctrm =
358
+ − 212
let
562
+ − 213
val cty = Thm.ctyp_of_cterm ctrm
358
+ − 214
in
562
+ − 215
Thm.instantiate' [SOME cty] [NONE, SOME ctrm] @{thm spec}
565
+ − 216
end\<close>
164
+ − 217
565
+ − 218
text \<open>
562
+ − 219
This helper function uses the function @{ML_ind instantiate' in Thm}
565
+ − 220
and instantiates the \<open>?x\<close> in the theorem @{thm spec} with a given
323
+ − 221
@{ML_type cterm}. We call this helper function in the following
+ − 222
tactic.\label{fun:instspectac}.
565
+ − 223
\<close>
164
+ − 224
565
+ − 225
ML %grayML\<open>fun inst_spec_tac ctxt ctrms =
+ − 226
EVERY' (map (dresolve_tac ctxt o single o inst_spec) ctrms)\<close>
183
+ − 227
565
+ − 228
text \<open>
212
+ − 229
This tactic expects a list of @{ML_type cterm}s. It allows us in the
+ − 230
proof below to instantiate the three quantifiers in the assumption.
565
+ − 231
\<close>
183
+ − 232
184
+ − 233
lemma
224
+ − 234
fixes P::"nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
+ − 235
shows "\<forall>x y z. P x y z \<Longrightarrow> True"
565
+ − 236
apply (tactic \<open>
+ − 237
inst_spec_tac @{context} [@{cterm "a::nat"},@{cterm "b::nat"},@{cterm "c::nat"}] 1\<close>)
+ − 238
txt \<open>
183
+ − 239
We obtain the goal state
179
+ − 240
+ − 241
\begin{minipage}{\textwidth}
+ − 242
@{subgoals}
565
+ − 243
\end{minipage}\<close>
164
+ − 244
(*<*)oops(*>*)
+ − 245
565
+ − 246
text \<open>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 247
The complete tactic for proving the induction principles can now
183
+ − 248
be implemented as follows:
565
+ − 249
\<close>
163
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 250
565
+ − 251
ML %linenosgray\<open>fun ind_tac ctxt defs prem insts =
552
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 252
EVERY1 [Object_Logic.full_atomize_tac ctxt,
208
+ − 253
cut_facts_tac prem,
552
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 254
rewrite_goal_tac ctxt defs,
562
+ − 255
inst_spec_tac ctxt insts,
565
+ − 256
assume_tac ctxt]\<close>
164
+ − 257
565
+ − 258
text \<open>
215
+ − 259
We have to give it as arguments the definitions, the premise (a list of
565
+ − 260
formulae) and the instantiations. The premise is \<open>even n\<close> in lemma
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 261
@{thm [source] manual_ind_prin_even} shown above; in our code it will always be a list
215
+ − 262
consisting of a single formula. Compare this tactic with the manual proof
+ − 263
for the lemma @{thm [source] manual_ind_prin_even}: as you can see there is
+ − 264
almost a one-to-one correspondence between the \isacommand{apply}-script and
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 265
the @{ML ind_tac}. We first rewrite the goal to use only object connectives (Line 2),
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 266
"cut in" the premise (Line 3), unfold the definitions (Line 4), instantiate
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 267
the assumptions of the goal (Line 5) and then conclude with @{ML assume_tac}.
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 268
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 269
Two testcases for this tactic are:
565
+ − 270
\<close>
183
+ − 271
210
+ − 272
lemma automatic_ind_prin_even:
208
+ − 273
assumes prem: "even z"
+ − 274
shows "P 0 \<Longrightarrow> (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P z"
565
+ − 275
by (tactic \<open>ind_tac @{context} eo_defs @{thms prem}
+ − 276
[@{cterm "P::nat\<Rightarrow>bool"}, @{cterm "Q::nat\<Rightarrow>bool"}]\<close>)
210
+ − 277
+ − 278
lemma automatic_ind_prin_fresh:
+ − 279
assumes prem: "fresh z za"
+ − 280
shows "(\<And>a b. a \<noteq> b \<Longrightarrow> P a (Var b)) \<Longrightarrow>
+ − 281
(\<And>a t s. \<lbrakk>P a t; P a s\<rbrakk> \<Longrightarrow> P a (App t s)) \<Longrightarrow>
+ − 282
(\<And>a t. P a (Lam a t)) \<Longrightarrow>
+ − 283
(\<And>a b t. \<lbrakk>a \<noteq> b; P a t\<rbrakk> \<Longrightarrow> P a (Lam b t)) \<Longrightarrow> P z za"
565
+ − 284
by (tactic \<open>ind_tac @{context} @{thms fresh_def} @{thms prem}
+ − 285
[@{cterm "P::string\<Rightarrow>trm\<Rightarrow>bool"}]\<close>)
210
+ − 286
164
+ − 287
565
+ − 288
text \<open>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 289
While the tactic for proving the induction principles is relatively simple,
237
+ − 290
it will be a bit more work to construct the goals from the introduction rules
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 291
the user provides. Therefore let us have a closer look at the first
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 292
proved theorem:
209
+ − 293
+ − 294
\begin{isabelle}
210
+ − 295
\isacommand{thm}~@{thm [source] automatic_ind_prin_even}\\
565
+ − 296
\<open>> \<close>~@{thm automatic_ind_prin_even}
209
+ − 297
\end{isabelle}
+ − 298
565
+ − 299
The variables \<open>z\<close>, \<open>P\<close> and \<open>Q\<close> are schematic
215
+ − 300
variables (since they are not quantified in the lemma). These
+ − 301
variables must be schematic, otherwise they cannot be instantiated
+ − 302
by the user. To generate these schematic variables we use a common trick
+ − 303
in Isabelle programming: we first declare them as \emph{free},
+ − 304
\emph{but fixed}, and then use the infrastructure to turn them into
+ − 305
schematic variables.
+ − 306
565
+ − 307
In general we have to construct for each predicate \<open>pred\<close> a goal
215
+ − 308
of the form
179
+ − 309
+ − 310
@{text [display]
208
+ − 311
"pred ?zs \<Longrightarrow> rules[preds := ?Ps] \<Longrightarrow> ?P ?zs"}
180
+ − 312
565
+ − 313
where the predicates \<open>preds\<close> are replaced in \<open>rules\<close> by new
+ − 314
distinct variables \<open>?Ps\<close>. We also need to generate fresh arguments
+ − 315
\<open>?zs\<close> for the predicate \<open>pred\<close> and the \<open>?P\<close> in
215
+ − 316
the conclusion.
190
+ − 317
565
+ − 318
We generate these goals in two steps. The first function, named \<open>prove_ind\<close>,
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 319
expects that the introduction rules are already appropriately substituted. The argument
565
+ − 320
\<open>srules\<close> stands for these substituted rules; \<open>cnewpreds\<close> are
+ − 321
the certified terms coresponding to the variables \<open>?Ps\<close>; \<open>pred\<close> is the predicate for which we prove the induction principle;
+ − 322
\<open>newpred\<close> is its replacement and \<open>arg_tys\<close> are the argument
208
+ − 323
types of this predicate.
565
+ − 324
\<close>
165
+ − 325
565
+ − 326
ML %linenosgray\<open>fun prove_ind lthy defs srules cnewpreds ((pred, newpred), arg_tys) =
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 327
let
190
+ − 328
val zs = replicate (length arg_tys) "z"
179
+ − 329
val (newargnames, lthy') = Variable.variant_fixes zs lthy;
190
+ − 330
val newargs = map Free (newargnames ~~ arg_tys)
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 331
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 332
val prem = HOLogic.mk_Trueprop (list_comb (pred, newargs))
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 333
val goal = Logic.list_implies
183
+ − 334
(srules, HOLogic.mk_Trueprop (list_comb (newpred, newargs)))
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 335
in
179
+ − 336
Goal.prove lthy' [] [prem] goal
552
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 337
(fn {prems, context, ...} => ind_tac context defs prems cnewpreds)
475
+ − 338
|> singleton (Proof_Context.export lthy' lthy)
565
+ − 339
end\<close>
179
+ − 340
565
+ − 341
text \<open>
+ − 342
In Line 3 we produce names \<open>zs\<close> for each type in the
183
+ − 343
argument type list. Line 4 makes these names unique and declares them as
565
+ − 344
free, but fixed, variables in the local theory \<open>lthy'\<close>.
212
+ − 345
That means they are not schematic variables (yet).
210
+ − 346
In Line 5 we construct the terms corresponding to these variables.
208
+ − 347
The variables are applied to the predicate in Line 7 (this corresponds
565
+ − 348
to the first premise \<open>pred zs\<close> of the induction principle).
+ − 349
In Line 8 and 9, we first construct the term \<open>P zs\<close>
212
+ − 350
and then add the (substituted) introduction rules as preconditions. In
+ − 351
case that no introduction rules are given, the conclusion of this
+ − 352
implication needs to be wrapped inside a @{term Trueprop}, otherwise
+ − 353
the Isabelle's goal mechanism will fail.\footnote{FIXME: check with
+ − 354
Stefan...is this so?}
180
+ − 355
316
+ − 356
In Line 11 we set up the goal to be proved using the function @{ML_ind
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 357
prove in Goal}; in the next line we call the tactic for proving the
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 358
induction principle. As mentioned before, this tactic expects the
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 359
definitions, the premise and the (certified) predicates with which the
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 360
introduction rules have been substituted. The code in these two lines will
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 361
return a theorem. However, it is a theorem proved inside the local theory
565
+ − 362
\<open>lthy'\<close>, where the variables \<open>zs\<close> are free, but fixed (see
+ − 363
Line 4). By exporting this theorem from \<open>lthy'\<close> (which contains the
+ − 364
\<open>zs\<close> as free variables) to \<open>lthy\<close> (which does not), we
+ − 365
obtain the desired schematic variables \<open>?zs\<close>. A testcase for this
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 366
function is
565
+ − 367
\<close>
180
+ − 368
565
+ − 369
local_setup %gray \<open>fn lthy =>
190
+ − 370
let
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 371
val newpreds = [@{term "P::nat \<Rightarrow> bool"}, @{term "Q::nat \<Rightarrow> bool"}]
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 372
val cnewpreds = [@{cterm "P::nat \<Rightarrow> bool"}, @{cterm "Q::nat \<Rightarrow> bool"}]
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 373
val newpred = @{term "P::nat \<Rightarrow> bool"}
210
+ − 374
val srules = map (subst_free (eo_preds ~~ newpreds)) eo_rules
190
+ − 375
val intro =
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 376
prove_ind lthy eo_defs srules cnewpreds ((e_pred, newpred), e_arg_tys)
190
+ − 377
in
440
+ − 378
pwriteln (pretty_thm lthy intro); lthy
565
+ − 379
end\<close>
184
+ − 380
565
+ − 381
text \<open>
210
+ − 382
This prints out the theorem:
190
+ − 383
+ − 384
@{text [display]
+ − 385
" \<lbrakk>even ?z; P 0; \<And>n. Q n \<Longrightarrow> P (Suc n); \<And>n. P n \<Longrightarrow> Q (Suc n)\<rbrakk> \<Longrightarrow> P ?z"}
+ − 386
565
+ − 387
The export from \<open>lthy'\<close> to \<open>lthy\<close> in Line 13 above
+ − 388
has correctly turned the free, but fixed, \<open>z\<close> into a schematic
+ − 389
variable \<open>?z\<close>; the variables \<open>P\<close> and \<open>Q\<close> are not yet
208
+ − 390
schematic.
190
+ − 391
+ − 392
We still have to produce the new predicates with which the introduction
210
+ − 393
rules are substituted and iterate @{ML prove_ind} over all
565
+ − 394
predicates. This is what the second function, named \<open>inds\<close> does.
+ − 395
\<close>
165
+ − 396
565
+ − 397
ML %linenosgray\<open>fun inds rules defs preds arg_tyss lthy =
164
+ − 398
let
+ − 399
val Ps = replicate (length preds) "P"
183
+ − 400
val (newprednames, lthy') = Variable.variant_fixes Ps lthy
164
+ − 401
184
+ − 402
val tyss' = map (fn tys => tys ---> HOLogic.boolT) arg_tyss
165
+ − 403
val newpreds = map Free (newprednames ~~ tyss')
562
+ − 404
val cnewpreds = map (Thm.cterm_of lthy') newpreds
184
+ − 405
val srules = map (subst_free (preds ~~ newpreds)) rules
164
+ − 406
+ − 407
in
210
+ − 408
map (prove_ind lthy' defs srules cnewpreds)
184
+ − 409
(preds ~~ newpreds ~~ arg_tyss)
475
+ − 410
|> Proof_Context.export lthy' lthy
565
+ − 411
end\<close>
165
+ − 412
565
+ − 413
text \<open>
208
+ − 414
In Line 3, we generate a string @{text [quotes] "P"} for each predicate.
184
+ − 415
In Line 4, we use the same trick as in the previous function, that is making the
565
+ − 416
\<open>Ps\<close> fresh and declaring them as free, but fixed, in
+ − 417
the new local theory \<open>lthy'\<close>. In Line 6, we construct the types of these new predicates
190
+ − 418
using the given argument types. Next we turn them into terms and subsequently
562
+ − 419
certify them (Line 7 and 8). We can now produce the substituted introduction rules
+ − 420
(Line 9) using the function @{ML_ind subst_free in Term}. Line 12 and 13 just iterate
190
+ − 421
the proofs for all predicates.
184
+ − 422
From this we obtain a list of theorems. Finally we need to export the
565
+ − 423
fixed variables \<open>Ps\<close> to obtain the schematic variables \<open>?Ps\<close>
562
+ − 424
(Line 14).
184
+ − 425
+ − 426
A testcase for this function is
565
+ − 427
\<close>
184
+ − 428
565
+ − 429
local_setup %gray \<open>fn lthy =>
184
+ − 430
let
210
+ − 431
val ind_thms = inds eo_rules eo_defs eo_preds eo_arg_tyss lthy
184
+ − 432
in
440
+ − 433
pwriteln (pretty_thms lthy ind_thms); lthy
565
+ − 434
end\<close>
165
+ − 435
176
+ − 436
565
+ − 437
text \<open>
184
+ − 438
which prints out
+ − 439
+ − 440
@{text [display]
210
+ − 441
"even ?z \<Longrightarrow> ?P1 0 \<Longrightarrow>
+ − 442
(\<And>m. ?Pa1 m \<Longrightarrow> ?P1 (Suc m)) \<Longrightarrow> (\<And>m. ?P1 m \<Longrightarrow> ?Pa1 (Suc m)) \<Longrightarrow> ?P1 ?z,
+ − 443
odd ?z \<Longrightarrow> ?P1 0 \<Longrightarrow>
+ − 444
(\<And>m. ?Pa1 m \<Longrightarrow> ?P1 (Suc m)) \<Longrightarrow> (\<And>m. ?P1 m \<Longrightarrow> ?Pa1 (Suc m)) \<Longrightarrow> ?Pa1 ?z"}
184
+ − 445
565
+ − 446
Note that now both, the \<open>?Ps\<close> and the \<open>?zs\<close>, are schematic
210
+ − 447
variables. The numbers attached to these variables have been introduced by
+ − 448
the pretty-printer and are \emph{not} important for the user.
184
+ − 449
210
+ − 450
This completes the code for the induction principles. The final peice
+ − 451
of reasoning infrastructure we need are the introduction rules.
565
+ − 452
\<close>
208
+ − 453
565
+ − 454
subsection \<open>Introduction Rules\<close>
208
+ − 455
565
+ − 456
text \<open>
212
+ − 457
Constructing the goals for the introduction rules is easy: they
+ − 458
are just the rules given by the user. However, their proofs are
+ − 459
quite a bit more involved than the ones for the induction principles.
+ − 460
To explain the general method, our running example will be
+ − 461
the introduction rule
208
+ − 462
212
+ − 463
\begin{isabelle}
+ − 464
@{prop "\<And>a b t. \<lbrakk>a \<noteq> b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"}
+ − 465
\end{isabelle}
+ − 466
+ − 467
about freshness for lambdas. In order to ease somewhat
+ − 468
our work here, we use the following two helper functions.
565
+ − 469
\<close>
184
+ − 470
565
+ − 471
ML %grayML\<open>val all_elims = fold (fn ct => fn th => th RS inst_spec ct)
+ − 472
val imp_elims = fold (fn th => fn th' => [th', th] MRS @{thm mp})\<close>
165
+ − 473
565
+ − 474
text \<open>
212
+ − 475
To see what these functions do, let us suppose we have the following three
190
+ − 476
theorems.
565
+ − 477
\<close>
190
+ − 478
+ − 479
lemma all_elims_test:
224
+ − 480
fixes P::"nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
+ − 481
shows "\<forall>x y z. P x y z" sorry
190
+ − 482
+ − 483
lemma imp_elims_test:
224
+ − 484
shows "A \<longrightarrow> B \<longrightarrow> C" sorry
190
+ − 485
+ − 486
lemma imp_elims_test':
224
+ − 487
shows "A" "B" sorry
190
+ − 488
565
+ − 489
text \<open>
190
+ − 490
The function @{ML all_elims} takes a list of (certified) terms and instantiates
+ − 491
theorems of the form @{thm [source] all_elims_test}. For example we can instantiate
210
+ − 492
the quantifiers in this theorem with @{term a}, @{term b} and @{term c} as follows:
190
+ − 493
572
+ − 494
@{ML_response [display, gray]
569
+ − 495
\<open>let
+ − 496
val ctrms = [@{cterm "a::nat"}, @{cterm "b::nat"}, @{cterm "c::nat"}]
190
+ − 497
val new_thm = all_elims ctrms @{thm all_elims_test}
+ − 498
in
440
+ − 499
pwriteln (pretty_thm_no_vars @{context} new_thm)
569
+ − 500
end\<close>
+ − 501
\<open>P a b c\<close>}
190
+ − 502
215
+ − 503
Note the difference with @{ML inst_spec_tac} from Page~\pageref{fun:instspectac}:
+ − 504
@{ML inst_spec_tac} is a tactic which operates on a goal state; in contrast
+ − 505
@{ML all_elims} operates on theorems.
+ − 506
190
+ − 507
Similarly, the function @{ML imp_elims} eliminates preconditions from implications.
565
+ − 508
For example we can eliminate the preconditions \<open>A\<close> and \<open>B\<close> from
210
+ − 509
@{thm [source] imp_elims_test}:
190
+ − 510
572
+ − 511
@{ML_response [display, gray]
569
+ − 512
\<open>let
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 513
val res = imp_elims @{thms imp_elims_test'} @{thm imp_elims_test}
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 514
in
440
+ − 515
pwriteln (pretty_thm_no_vars @{context} res)
569
+ − 516
end\<close>
572
+ − 517
\<open>Q\<close>}
190
+ − 518
212
+ − 519
Now we set up the proof for the introduction rule as follows:
565
+ − 520
\<close>
190
+ − 521
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 522
lemma fresh_Lam:
224
+ − 523
shows "\<And>a b t. \<lbrakk>a \<noteq> b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
210
+ − 524
(*<*)oops(*>*)
+ − 525
565
+ − 526
text \<open>
212
+ − 527
The first step in the proof will be to expand the definitions of freshness
210
+ − 528
and then introduce quantifiers and implications. For this we
+ − 529
will use the tactic
565
+ − 530
\<close>
210
+ − 531
565
+ − 532
ML %linenosgray\<open>fun expand_tac ctxt defs =
552
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 533
Object_Logic.rulify_tac ctxt 1
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 534
THEN rewrite_goal_tac ctxt defs 1
565
+ − 535
THEN (REPEAT (resolve_tac ctxt [@{thm allI}, @{thm impI}] 1))\<close>
210
+ − 536
565
+ − 537
text \<open>
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 538
The function in Line 2 ``rulifies'' the lemma.\footnote{FIXME: explain this better}
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 539
This will turn out to
565
+ − 540
be important later on. Applying this tactic in our proof of \<open>fresh_Lem\<close>
+ − 541
\<close>
210
+ − 542
+ − 543
(*<*)
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 544
lemma fresh_Lam:
224
+ − 545
shows "\<And>a b t. \<lbrakk>a \<noteq> b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
210
+ − 546
(*>*)
565
+ − 547
apply(tactic \<open>expand_tac @{context} @{thms fresh_def}\<close>)
209
+ − 548
565
+ − 549
txt \<open>
215
+ − 550
gives us the goal state
210
+ − 551
209
+ − 552
\begin{isabelle}
210
+ − 553
@{subgoals [display]}
209
+ − 554
\end{isabelle}
210
+ − 555
565
+ − 556
As you can see, there are parameters (namely \<open>a\<close>, \<open>b\<close> and
+ − 557
\<open>t\<close>) which come from the introduction rule and parameters (in the
+ − 558
case above only \<open>fresh\<close>) which come from the universal
215
+ − 559
quantification in the definition @{term "fresh a (App t s)"}. Similarly,
+ − 560
there are assumptions that come from the premises of the rule (namely the
+ − 561
first two) and assumptions from the definition of the predicate (assumption
+ − 562
three to six). We need to treat these parameters and assumptions
565
+ − 563
differently. In the code below we will therefore separate them into \<open>params1\<close> and \<open>params2\<close>, respectively \<open>prems1\<close> and \<open>prems2\<close>. To do this separation, it is best to open a subproof with the
369
+ − 564
tactic @{ML_ind SUBPROOF in Subgoal}, since this tactic provides us with the parameters (as
215
+ − 565
list of @{ML_type cterm}s) and the assumptions (as list of @{ML_type thm}s).
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 566
The problem with @{ML SUBPROOF}, however, is that it always expects us to
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 567
completely discharge the goal (see Section~\ref{sec:simpletacs}). This is
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 568
a bit inconvenient for our gradual explanation of the proof here. Therefore
316
+ − 569
we use first the function @{ML_ind FOCUS in Subgoal}, which does s
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 570
ame as @{ML SUBPROOF}
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 571
but does not require us to completely discharge the goal.
565
+ − 572
\<close>
210
+ − 573
(*<*)oops(*>*)
565
+ − 574
text_raw \<open>
210
+ − 575
\begin{figure}[t]
+ − 576
\begin{minipage}{\textwidth}
+ − 577
\begin{isabelle}
565
+ − 578
\<close>
+ − 579
ML %grayML\<open>fun chop_print params1 params2 prems1 prems2 ctxt =
210
+ − 580
let
574
+ − 581
val pps = [Pretty.big_list "Params1 from the rule:"
+ − 582
(map (pretty_cterm ctxt) params1),
+ − 583
Pretty.big_list "Params2 from the predicate:"
+ − 584
(map (pretty_cterm ctxt) params2),
+ − 585
Pretty.big_list "Prems1 from the rule:"
+ − 586
(map (pretty_thm ctxt) prems1),
+ − 587
Pretty.big_list "Prems2 from the predicate:"
+ − 588
(map (pretty_thm ctxt) prems2)]
210
+ − 589
in
440
+ − 590
pps |> Pretty.chunks
448
+ − 591
|> Pretty.string_of
+ − 592
|> tracing
565
+ − 593
end\<close>
448
+ − 594
565
+ − 595
text_raw\<open>
210
+ − 596
\end{isabelle}
+ − 597
\end{minipage}
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 598
\caption{A helper function that prints out the parameters and premises that
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 599
need to be treated differently.\label{fig:chopprint}}
210
+ − 600
\end{figure}
565
+ − 601
\<close>
210
+ − 602
565
+ − 603
text \<open>
+ − 604
First we calculate the values for \<open>params1/2\<close> and \<open>prems1/2\<close>
+ − 605
from \<open>params\<close> and \<open>prems\<close>, respectively. To better see what is
212
+ − 606
going in our example, we will print out these values using the printing
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 607
function in Figure~\ref{fig:chopprint}. Since @{ML FOCUS in Subgoal} will
565
+ − 608
supply us the \<open>params\<close> and \<open>prems\<close> as lists, we can
369
+ − 609
separate them using the function @{ML_ind chop in Library}.
565
+ − 610
\<close>
210
+ − 611
565
+ − 612
ML %linenosgray\<open>fun chop_test_tac preds rules =
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 613
Subgoal.FOCUS (fn {params, prems, context, ...} =>
210
+ − 614
let
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 615
val cparams = map snd params
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 616
val (params1, params2) = chop (length cparams - length preds) cparams
210
+ − 617
val (prems1, prems2) = chop (length prems - length rules) prems
+ − 618
in
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 619
chop_print params1 params2 prems1 prems2 context; all_tac
565
+ − 620
end)\<close>
210
+ − 621
565
+ − 622
text \<open>
212
+ − 623
For the separation we can rely on the fact that Isabelle deterministically
+ − 624
produces parameters and premises in a goal state. The last parameters
+ − 625
that were introduced come from the quantifications in the definitions
+ − 626
(see the tactic @{ML expand_tac}).
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 627
Therefore we only have to subtract in Line 5 the number of predicates (in this
565
+ − 628
case only \<open>1\<close>) from the lenghts of all parameters. Similarly
+ − 629
with the \<open>prems\<close> in line 6: the last premises in the goal state come from
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 630
unfolding the definition of the predicate in the conclusion. So we can
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 631
just subtract the number of rules from the number of all premises.
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 632
To check our calculations we print them out in Line 8 using the
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 633
function @{ML chop_print} from Figure~\ref{fig:chopprint} and then
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 634
just do nothing, that is @{ML all_tac}. Applying this tactic in our example
565
+ − 635
\<close>
209
+ − 636
210
+ − 637
(*<*)
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 638
lemma fresh_Lam:
224
+ − 639
shows "\<And>a b t. \<lbrakk>a \<noteq> b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
565
+ − 640
apply(tactic \<open>expand_tac @{context} @{thms fresh_def}\<close>)
210
+ − 641
(*>*)
565
+ − 642
apply(tactic \<open>chop_test_tac [fresh_pred] fresh_rules @{context} 1\<close>)
210
+ − 643
(*<*)oops(*>*)
+ − 644
565
+ − 645
text \<open>
210
+ − 646
gives
209
+ − 647
210
+ − 648
\begin{isabelle}
565
+ − 649
\<open>Params1 from the rule:\<close>\\
+ − 650
\<open>a, b, t\<close>\\
+ − 651
\<open>Params2 from the predicate:\<close>\\
+ − 652
\<open>fresh\<close>\\
+ − 653
\<open>Prems1 from the rule:\<close>\\
210
+ − 654
@{term "a \<noteq> b"}\\
574
+ − 655
@{text "\<forall>fresh."}\\
+ − 656
@{text "(\<forall>a b. a \<noteq> b \<longrightarrow> fresh a (Var b)) \<longrightarrow>"}\\
+ − 657
@{text "(\<forall>a t s. fresh a t \<longrightarrow> fresh a s \<longrightarrow> fresh a (App t s)) \<longrightarrow>"}\\
+ − 658
@{text "(\<forall>a t. fresh a (Lam a t)) \<longrightarrow> "}\\
+ − 659
@{text "(\<forall>a b t. a \<noteq> b \<longrightarrow> fresh a t \<longrightarrow> fresh a (Lam b t)) \<longrightarrow> fresh a t"}\\
565
+ − 660
\<open>Prems2 from the predicate:\<close>\\
210
+ − 661
@{term "\<forall>a b. a \<noteq> b \<longrightarrow> fresh a (Var b)"}\\
+ − 662
@{term "\<forall>a t s. fresh a t \<longrightarrow> fresh a s \<longrightarrow> fresh a (App t s)"}\\
+ − 663
@{term "\<forall>a t. fresh a (Lam a t)"}\\
+ − 664
@{term "\<forall>a b t. a \<noteq> b \<longrightarrow> fresh a t \<longrightarrow> fresh a (Lam b t)"}
+ − 665
\end{isabelle}
208
+ − 666
192
+ − 667
565
+ − 668
We now have to select from \<open>prems2\<close> the premise
210
+ − 669
that corresponds to the introduction rule we prove, namely:
+ − 670
212
+ − 671
@{term [display] "\<forall>a b t. a \<noteq> b \<longrightarrow> fresh a t \<longrightarrow> fresh a (Lam a t)"}
210
+ − 672
562
+ − 673
To use this premise with @{ML resolve_tac}, we need to instantiate its
565
+ − 674
quantifiers (with \<open>params1\<close>) and transform it into rule
418
+ − 675
format (using @{ML_ind rulify in Object_Logic}). So we can modify the
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 676
code as follows:
565
+ − 677
\<close>
210
+ − 678
565
+ − 679
ML %linenosgray\<open>fun apply_prem_tac i preds rules =
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 680
Subgoal.FOCUS (fn {params, prems, context, ...} =>
210
+ − 681
let
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 682
val cparams = map snd params
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 683
val (params1, params2) = chop (length cparams - length preds) cparams
210
+ − 684
val (prems1, prems2) = chop (length prems - length rules) prems
+ − 685
in
574
+ − 686
resolve_tac context
+ − 687
[Object_Logic.rulify context (all_elims params1 (nth prems2 i))] 1
565
+ − 688
end)\<close>
210
+ − 689
565
+ − 690
text \<open>
+ − 691
The argument \<open>i\<close> corresponds to the number of the
215
+ − 692
introduction we want to prove. We will later on let it range
565
+ − 693
from \<open>0\<close> to the number of \<open>rules - 1\<close>.
+ − 694
Below we apply this function with \<open>3\<close>, since
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 695
we are proving the fourth introduction rule.
565
+ − 696
\<close>
210
+ − 697
+ − 698
(*<*)
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 699
lemma fresh_Lam:
224
+ − 700
shows "\<And>a b t. \<lbrakk>a \<noteq> b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
565
+ − 701
apply(tactic \<open>expand_tac @{context} @{thms fresh_def}\<close>)
210
+ − 702
(*>*)
565
+ − 703
apply(tactic \<open>apply_prem_tac 3 [fresh_pred] fresh_rules @{context} 1\<close>)
210
+ − 704
(*<*)oops(*>*)
+ − 705
565
+ − 706
text \<open>
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 707
The goal state we obtain is:
210
+ − 708
+ − 709
\begin{isabelle}
565
+ − 710
\<open>1.\<close>~\<open>\<dots> \<Longrightarrow> \<close>~@{prop "a \<noteq> b"}\\
+ − 711
\<open>2.\<close>~\<open>\<dots> \<Longrightarrow> \<close>~@{prop "fresh a t"}
210
+ − 712
\end{isabelle}
+ − 713
215
+ − 714
As expected there are two subgoals, where the first comes from the
212
+ − 715
non-recursive premise of the introduction rule and the second comes
215
+ − 716
from the recursive one. The first goal can be solved immediately
565
+ − 717
by \<open>prems1\<close>. The second needs more work. It can be solved
+ − 718
with the other premise in \<open>prems1\<close>, namely
212
+ − 719
210
+ − 720
+ − 721
@{term [break,display]
+ − 722
"\<forall>fresh.
+ − 723
(\<forall>a b. a \<noteq> b \<longrightarrow> fresh a (Var b)) \<longrightarrow>
+ − 724
(\<forall>a t s. fresh a t \<longrightarrow> fresh a s \<longrightarrow> fresh a (App t s)) \<longrightarrow>
+ − 725
(\<forall>a t. fresh a (Lam a t)) \<longrightarrow>
+ − 726
(\<forall>a b t. a \<noteq> b \<longrightarrow> fresh a t \<longrightarrow> fresh a (Lam b t)) \<longrightarrow> fresh a t"}
+ − 727
+ − 728
but we have to instantiate it appropriately. These instantiations
565
+ − 729
come from \<open>params1\<close> and \<open>prems2\<close>. We can determine
210
+ − 730
whether we are in the simple or complicated case by checking whether
565
+ − 731
the topmost connective is an \<open>\<forall>\<close>. The premises in the simple
212
+ − 732
case cannot have such a quantification, since the first step
569
+ − 733
of @{ML \<open>expand_tac\<close>} was to ``rulify'' the lemma.
565
+ − 734
The premise of the complicated case must have at least one \<open>\<forall>\<close>
+ − 735
coming from the quantification over the \<open>preds\<close>. So
210
+ − 736
we can implement the following function
565
+ − 737
\<close>
210
+ − 738
565
+ − 739
ML %grayML\<open>fun prepare_prem ctxt params2 prems2 prem =
562
+ − 740
resolve_tac ctxt [case Thm.prop_of prem of
165
+ − 741
_ $ (Const (@{const_name All}, _) $ _) =>
210
+ − 742
prem |> all_elims params2
+ − 743
|> imp_elims prems2
565
+ − 744
| _ => prem]\<close>
210
+ − 745
565
+ − 746
text \<open>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 747
which either applies the premise outright (the default case) or if
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 748
it has an outermost universial quantification, instantiates it first
565
+ − 749
with \<open>params1\<close> and then \<open>prems1\<close>. The following
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 750
tactic will therefore prove the lemma completely.
565
+ − 751
\<close>
210
+ − 752
565
+ − 753
ML %grayML\<open>fun prove_intro_tac i preds rules =
552
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 754
SUBPROOF (fn {params, prems, context, ...} =>
210
+ − 755
let
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 756
val cparams = map snd params
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 757
val (params1, params2) = chop (length cparams - length preds) cparams
210
+ − 758
val (prems1, prems2) = chop (length prems - length rules) prems
+ − 759
in
574
+ − 760
resolve_tac context
+ − 761
[Object_Logic.rulify context (all_elims params1 (nth prems2 i))] 1
562
+ − 762
THEN EVERY1 (map (prepare_prem context params2 prems2) prems1)
565
+ − 763
end)\<close>
210
+ − 764
565
+ − 765
text \<open>
299
d0b81d6e1b28
updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 766
Note that the tactic is now @{ML SUBPROOF}, not @{ML FOCUS in Subgoal} anymore.
215
+ − 767
The full proof of the introduction rule is as follows:
565
+ − 768
\<close>
210
+ − 769
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 770
lemma fresh_Lam:
224
+ − 771
shows "\<And>a b t. \<lbrakk>a \<noteq> b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
565
+ − 772
apply(tactic \<open>expand_tac @{context} @{thms fresh_def}\<close>)
+ − 773
apply(tactic \<open>prove_intro_tac 3 [fresh_pred] fresh_rules @{context} 1\<close>)
210
+ − 774
done
+ − 775
565
+ − 776
text \<open>
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 777
Phew!\ldots
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 778
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 779
Unfortunately, not everything is done yet. If you look closely
212
+ − 780
at the general principle outlined for the introduction rules in
+ − 781
Section~\ref{sec:nutshell}, we have not yet dealt with the case where
+ − 782
recursive premises have preconditions. The introduction rule
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 783
of the accessible part is such a rule.
565
+ − 784
\<close>
210
+ − 785
448
+ − 786
210
+ − 787
lemma accpartI:
224
+ − 788
shows "\<And>R x. (\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
565
+ − 789
apply(tactic \<open>expand_tac @{context} @{thms accpart_def}\<close>)
+ − 790
apply(tactic \<open>chop_test_tac [acc_pred] acc_rules @{context} 1\<close>)
+ − 791
apply(tactic \<open>apply_prem_tac 0 [acc_pred] acc_rules @{context} 1\<close>)
210
+ − 792
565
+ − 793
txt \<open>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 794
Here @{ML chop_test_tac} prints out the following
565
+ − 795
values for \<open>params1/2\<close> and \<open>prems1/2\<close>
210
+ − 796
+ − 797
\begin{isabelle}
565
+ − 798
\<open>Params1 from the rule:\<close>\\
+ − 799
\<open>x\<close>\\
+ − 800
\<open>Params2 from the predicate:\<close>\\
+ − 801
\<open>P\<close>\\
+ − 802
\<open>Prems1 from the rule:\<close>\\
+ − 803
\<open>R ?y x \<Longrightarrow> \<forall>P. (\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> P ?y\<close>\\
+ − 804
\<open>Prems2 from the predicate:\<close>\\
210
+ − 805
@{term "\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x"}\\
+ − 806
\end{isabelle}
+ − 807
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 808
and after application of the introduction rule
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 809
using @{ML apply_prem_tac}, we are in the goal state
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 810
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 811
\begin{isabelle}
565
+ − 812
\<open>1.\<close>~@{term "\<And>y. R y x \<Longrightarrow> P y"}
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 813
\end{isabelle}
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 814
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 815
565
+ − 816
\<close>(*<*)oops(*>*)
210
+ − 817
565
+ − 818
text \<open>
212
+ − 819
In order to make progress, we have to use the precondition
565
+ − 820
\<open>R y x\<close> (in general there can be many of them). The best way
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 821
to get a handle on these preconditions is to open up another subproof,
565
+ − 822
since the preconditions will then be bound to \<open>prems\<close>. Therfore we
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 823
modify the function @{ML prepare_prem} as follows
565
+ − 824
\<close>
210
+ − 825
565
+ − 826
ML %linenosgray\<open>fun prepare_prem params2 prems2 ctxt prem =
210
+ − 827
SUBPROOF (fn {prems, ...} =>
+ − 828
let
+ − 829
val prem' = prems MRS prem
+ − 830
in
562
+ − 831
resolve_tac ctxt [case Thm.prop_of prem' of
210
+ − 832
_ $ (Const (@{const_name All}, _) $ _) =>
+ − 833
prem' |> all_elims params2
+ − 834
|> imp_elims prems2
562
+ − 835
| _ => prem'] 1
565
+ − 836
end) ctxt\<close>
210
+ − 837
565
+ − 838
text \<open>
+ − 839
In Line 4 we use the \<open>prems\<close> from the @{ML SUBPROOF} and resolve
+ − 840
them with \<open>prem\<close>. In the simple cases, that is where the \<open>prem\<close>
+ − 841
comes from a non-recursive premise of the rule, \<open>prems\<close> will be
369
+ − 842
just the empty list and the function @{ML_ind MRS in Drule} does nothing. Similarly, in the
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 843
cases where the recursive premises of the rule do not have preconditions.
212
+ − 844
In case there are preconditions, then Line 4 discharges them. After
+ − 845
that we can proceed as before, i.e., check whether the outermost
565
+ − 846
connective is \<open>\<forall>\<close>.
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 847
212
+ − 848
The function @{ML prove_intro_tac} only needs to be changed so that it
575
+ − 849
gives the context to @{ML prepare_prem} (Line 10). The modified version
212
+ − 850
is below.
565
+ − 851
\<close>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 852
565
+ − 853
ML %linenosgray\<open>fun prove_intro_tac i preds rules =
210
+ − 854
SUBPROOF (fn {params, prems, context, ...} =>
+ − 855
let
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 856
val cparams = map snd params
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 857
val (params1, params2) = chop (length cparams - length preds) cparams
210
+ − 858
val (prems1, prems2) = chop (length prems - length rules) prems
+ − 859
in
574
+ − 860
resolve_tac context
+ − 861
[Object_Logic.rulify context (all_elims params1 (nth prems2 i))] 1
210
+ − 862
THEN EVERY1 (map (prepare_prem params2 prems2 context) prems1)
565
+ − 863
end)\<close>
210
+ − 864
565
+ − 865
text \<open>
212
+ − 866
With these two functions we can now also prove the introduction
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 867
rule for the accessible part.
565
+ − 868
\<close>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 869
210
+ − 870
lemma accpartI:
224
+ − 871
shows "\<And>R x. (\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
565
+ − 872
apply(tactic \<open>expand_tac @{context} @{thms accpart_def}\<close>)
+ − 873
apply(tactic \<open>prove_intro_tac 0 [acc_pred] acc_rules @{context} 1\<close>)
210
+ − 874
done
+ − 875
565
+ − 876
text \<open>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 877
Finally we need two functions that string everything together. The first
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 878
function is the tactic that performs the proofs.
565
+ − 879
\<close>
190
+ − 880
565
+ − 881
ML %linenosgray\<open>fun intro_tac defs rules preds i ctxt =
552
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 882
EVERY1 [Object_Logic.rulify_tac ctxt,
Christian Urban <christian dot urban at kcl dot ac dot uk>
diff
changeset
+ − 883
rewrite_goal_tac ctxt defs,
562
+ − 884
REPEAT o (resolve_tac ctxt [@{thm allI}, @{thm impI}]),
565
+ − 885
prove_intro_tac i preds rules ctxt]\<close>
165
+ − 886
565
+ − 887
text \<open>
215
+ − 888
Lines 2 to 4 in this tactic correspond to the function @{ML expand_tac}.
+ − 889
Some testcases for this tactic are:
565
+ − 890
\<close>
190
+ − 891
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 892
lemma even0_intro:
224
+ − 893
shows "even 0"
565
+ − 894
by (tactic \<open>intro_tac eo_defs eo_rules eo_preds 0 @{context}\<close>)
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 895
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 896
lemma evenS_intro:
224
+ − 897
shows "\<And>m. odd m \<Longrightarrow> even (Suc m)"
565
+ − 898
by (tactic \<open>intro_tac eo_defs eo_rules eo_preds 1 @{context}\<close>)
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 899
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 900
lemma fresh_App:
224
+ − 901
shows "\<And>a t s. \<lbrakk>fresh a t; fresh a s\<rbrakk> \<Longrightarrow> fresh a (App t s)"
565
+ − 902
by (tactic \<open>
+ − 903
intro_tac @{thms fresh_def} fresh_rules [fresh_pred] 1 @{context}\<close>)
190
+ − 904
565
+ − 905
text \<open>
215
+ − 906
The second function sets up in Line 4 the goals to be proved (this is easy
212
+ − 907
for the introduction rules since they are exactly the rules
+ − 908
given by the user) and iterates @{ML intro_tac} over all
+ − 909
introduction rules.
565
+ − 910
\<close>
173
+ − 911
565
+ − 912
ML %linenosgray\<open>fun intros rules preds defs lthy =
165
+ − 913
let
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 914
fun intros_aux (i, goal) =
165
+ − 915
Goal.prove lthy [] [] goal
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 916
(fn {context, ...} => intro_tac defs rules preds i context)
165
+ − 917
in
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 918
map_index intros_aux rules
565
+ − 919
end\<close>
164
+ − 920
565
+ − 921
text \<open>
369
+ − 922
The iteration is done with the function @{ML_ind map_index in Library} since we
212
+ − 923
need the introduction rule together with its number (counted from
565
+ − 924
\<open>0\<close>). This completes the code for the functions deriving the
212
+ − 925
reasoning infrastructure. It remains to implement some administrative
+ − 926
code that strings everything together.
565
+ − 927
\<close>
212
+ − 928
565
+ − 929
subsection \<open>Administrative Functions\<close>
215
+ − 930
565
+ − 931
text \<open>
215
+ − 932
We have produced various theorems (definitions, induction principles and
+ − 933
introduction rules), but apart from the definitions, we have not yet
+ − 934
registered them with the theorem database. This is what the functions
394
+ − 935
@{ML_ind note in Local_Theory} does.
215
+ − 936
+ − 937
+ − 938
For convenience, we use the following
+ − 939
three wrappers this function:
565
+ − 940
\<close>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 941
565
+ − 942
ML %grayML\<open>fun note_many qname ((name, attrs), thms) =
394
+ − 943
Local_Theory.note ((Binding.qualify false qname name, attrs), thms)
215
+ − 944
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 945
fun note_single1 qname ((name, attrs), thm) =
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 946
note_many qname ((name, attrs), [thm])
176
+ − 947
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 948
fun note_single2 name attrs (qname, thm) =
565
+ − 949
note_many (Binding.name_of qname) ((name, attrs), [thm])\<close>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 950
565
+ − 951
text \<open>
+ − 952
The function that ``holds everything together'' is \<open>add_inductive\<close>.
+ − 953
Its arguments are the specification of the predicates \<open>pred_specs\<close>
+ − 954
and the introduction rules \<open>rule_spec\<close>.
+ − 955
\<close>
211
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 956
565
+ − 957
ML %linenosgray\<open>fun add_inductive pred_specs rule_specs lthy =
165
+ − 958
let
237
+ − 959
val mxs = map snd pred_specs
165
+ − 960
val pred_specs' = map fst pred_specs
+ − 961
val prednames = map fst pred_specs'
+ − 962
val preds = map (fn (p, ty) => Free (Binding.name_of p, ty)) pred_specs'
215
+ − 963
val tyss = map (binder_types o fastype_of) preds
163
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 964
215
+ − 965
val (namesattrs, rules) = split_list rule_specs
165
+ − 966
237
+ − 967
val (defs, lthy') = defns rules preds prednames mxs tyss lthy
+ − 968
val ind_prins = inds rules defs preds tyss lthy'
210
+ − 969
val intro_rules = intros rules preds defs lthy'
91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 970
165
+ − 971
val mut_name = space_implode "_" (map Binding.name_of prednames)
215
+ − 972
val case_names = map (Binding.name_of o fst) namesattrs
165
+ − 973
in
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 974
lthy' |> note_many mut_name ((@{binding "intros"}, []), intro_rules)
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 975
||>> note_many mut_name ((@{binding "inducts"}, []), ind_prins)
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 976
||>> fold_map (note_single1 mut_name) (namesattrs ~~ intro_rules)
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 977
||>> fold_map (note_single2 @{binding "induct"}
375
+ − 978
[Attrib.internal (K (Rule_Cases.case_names case_names)),
+ − 979
Attrib.internal (K (Rule_Cases.consumes 1)),
215
+ − 980
Attrib.internal (K (Induct.induct_pred ""))])
237
+ − 981
(prednames ~~ ind_prins)
215
+ − 982
|> snd
565
+ − 983
end\<close>
91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 984
565
+ − 985
text \<open>
215
+ − 986
In Line 3 the function extracts the syntax annotations from the predicates.
+ − 987
Lines 4 to 6 extract the names of the predicates and generate
+ − 988
the variables terms (with types) corresponding to the predicates.
+ − 989
Line 7 produces the argument types for each predicate.
+ − 990
+ − 991
Line 9 extracts the introduction rules from the specifications
565
+ − 992
and stores also in \<open>namesattrs\<close> the names and attributes the
215
+ − 993
user may have attached to these rules.
+ − 994
+ − 995
Line 11 produces the definitions and also registers the definitions
565
+ − 996
in the local theory \<open>lthy'\<close>. The next two lines produce
215
+ − 997
the induction principles and the introduction rules (all of them
565
+ − 998
as theorems). Both need the local theory \<open>lthy'\<close> in which
215
+ − 999
the definitions have been registered.
+ − 1000
+ − 1001
Lines 15 produces the name that is used to register the introduction
+ − 1002
rules. It is costum to collect all introduction rules under
565
+ − 1003
\<open>string.intros\<close>, whereby \<open>string\<close> stands for the
215
+ − 1004
@{text [quotes] "_"}-separated list of predicate names (for example
565
+ − 1005
\<open>even_odd\<close>. Also by custom, the case names in intuction
215
+ − 1006
proofs correspond to the names of the introduction rules. These
+ − 1007
are generated in Line 16.
+ − 1008
565
+ − 1009
Lines 18 and 19 now add to \<open>lthy'\<close> all the introduction rules
+ − 1010
und induction principles under the name \<open>mut_name.intros\<close> and
+ − 1011
\<open>mut_name.inducts\<close>, respectively (see previous paragraph).
237
+ − 1012
+ − 1013
Line 20 add further every introduction rule under its own name
215
+ − 1014
(given by the user).\footnote{FIXME: what happens if the user did not give
237
+ − 1015
any name.} Line 21 registers the induction principles. For this we have
375
+ − 1016
to use some specific attributes. The first @{ML_ind case_names in Rule_Cases}
215
+ − 1017
corresponds to the case names that are used by Isar to reference the proof
569
+ − 1018
obligations in the induction. The second @{ML \<open>consumes 1\<close> in Rule_Cases}
215
+ − 1019
indicates that the first premise of the induction principle (namely
+ − 1020
the predicate over which the induction proceeds) is eliminated.
+ − 1021
+ − 1022
This completes all the code and fits in with the ``front end'' described
237
+ − 1023
in Section~\ref{sec:interface}.\footnote{FIXME: Describe @{ML Induct.induct_pred}.
+ − 1024
Why the mut-name?
224
+ − 1025
What does @{ML Binding.qualify} do?}
565
+ − 1026
\<close>
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1027
(*<*)end(*>*)