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