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