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