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