91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 1
theory Ind_Code
176
+ − 2
imports "../Base" "../FirstSteps" Simple_Inductive_Package Ind_Prelims
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
164
+ − 5
section {* Code *}
+ − 6
+ − 7
text {*
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 8
@{text [display] "rule ::= \<And>xs. As \<Longrightarrow> (\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>* \<Longrightarrow> pred ts"}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 9
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 10
@{text [display] "orule ::= \<forall>xs. As \<longrightarrow> (\<forall>ys. Bs \<longrightarrow> pred ss)\<^isup>* \<longrightarrow> pred ts"}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 11
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 12
@{text [display] "def ::= 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
+ − 13
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 14
@{text [display] "ind ::= \<And>zs. pred zs \<Longrightarrow> rules[preds::=Ps] \<Longrightarrow> P zs"}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 15
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 16
@{text [display] "oind ::= \<forall>zs. pred zs \<longrightarrow> orules[preds::=Ps] \<longrightarrow> P zs"}
189
+ − 17
+ − 18
\underline{Induction proof}
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 19
189
+ − 20
After ``objectivication'' we have
+ − 21
@{text "pred zs"} and @{text "orules[preds::=Ps]"}; and have to show
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 22
@{text "P zs"}. Expanding @{text "pred zs"} gives @{text "\<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
+ − 23
Instantiating the @{text "preds"} with @{text "Ps"} gives
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 24
@{text "orules[preds::=Ps] \<longrightarrow> P zs"}. So we can conclude with @{text "P zs"}.
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 25
189
+ − 26
\underline{Intro proof}
+ − 27
+ − 28
Assume we want to prove the $i$th intro rule.
+ − 29
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 30
We have to show @{text "\<forall>xs. As \<longrightarrow> (\<forall>ys. Bs \<longrightarrow> pred ss)\<^isup>* \<longrightarrow> pred ts"};
189
+ − 31
expanding the defs, gives
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 32
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 33
@{text [display]
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 34
"\<forall>xs. As \<longrightarrow> (\<forall>ys. Bs \<longrightarrow> (\<forall>preds. orules \<longrightarrow> pred ss))\<^isup>* \<longrightarrow> (\<forall>preds. orules \<longrightarrow> pred ts"}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 35
189
+ − 36
applying as many allI and impI as possible
+ − 37
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 38
so we have @{text "As"}, @{text "(\<forall>ys. Bs \<longrightarrow> (\<forall>preds. orules \<longrightarrow> pred ss))\<^isup>*"},
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 39
@{text "orules"}; and have to show @{text "pred ts"}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 40
189
+ − 41
the $i$th @{text "orule"} is of the
+ − 42
form @{text "\<forall>xs. As \<longrightarrow> (\<forall>ys. Bs \<longrightarrow> pred ss)\<^isup>* \<longrightarrow> pred ts"}.
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 43
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 44
using the @{text "As"} we ????
164
+ − 45
*}
+ − 46
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 47
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 48
text {*
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 49
First we have to produce for each predicate its definitions of the form
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 50
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 51
@{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
+ − 52
183
+ − 53
In order to make definitions, we use the following wrapper for
+ − 54
@{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
+ − 55
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
+ − 56
*}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 57
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 58
ML %linenosgray{*fun make_defs ((predname, syn), trm) lthy =
164
+ − 59
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
+ − 60
val arg = ((predname, syn), (Attrib.empty_binding, trm))
176
+ − 61
val ((_, (_ , thm)), lthy') = LocalTheory.define Thm.internalK arg lthy
164
+ − 62
in
176
+ − 63
(thm, lthy')
164
+ − 64
end*}
+ − 65
+ − 66
text {*
183
+ − 67
It returns the definition (as a theorem) and the local theory in which this definition has
184
+ − 68
been made. In Line 4, @{ML internalK in Thm} is a flag attached to the
183
+ − 69
theorem (others possibilities are @{ML definitionK in Thm} and @{ML axiomK in Thm}).
176
+ − 70
These flags just classify theorems and have no significant meaning, except
183
+ − 71
for tools that, for example, find theorems in the theorem database. We also
179
+ − 72
use @{ML empty_binding in Attrib} in Line 3, since the definition does
183
+ − 73
not need to have any theorem attributes. A testcase for this function is
164
+ − 74
*}
+ − 75
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 76
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
+ − 77
let
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 78
val arg = ((@{binding "MyTrue"}, NoSyn), @{term True})
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 79
val (def, lthy') = make_defs arg lthy
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 80
in
179
+ − 81
warning (str_of_thm 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
+ − 82
end *}
164
+ − 83
+ − 84
text {*
184
+ − 85
which makes the definition @{prop "MyTrue \<equiv> True"} and then prints it out.
183
+ − 86
Since we are testing the function inside \isacommand{local\_setup}, i.e.~make
+ − 87
changes to the ambient theory, we can query the definition using the usual
+ − 88
command \isacommand{thm}:
179
+ − 89
+ − 90
\begin{isabelle}
+ − 91
\isacommand{thm}~@{text "MyTrue_def"}\\
+ − 92
@{text "> MyTrue \<equiv> True"}
+ − 93
\end{isabelle}
164
+ − 94
184
+ − 95
The next two functions construct the terms we need for the definitions for
+ − 96
our \isacommand{simple\_inductive} command. These
+ − 97
terms are of the form
179
+ − 98
+ − 99
@{text [display] "\<lambda>\<^raw:$zs$>. \<forall>preds. orules \<longrightarrow> pred \<^raw:$zs$>"}
+ − 100
183
+ − 101
The variables @{text "\<^raw:$zs$>"} need to be chosen so that they do not occur
184
+ − 102
in the @{text orules} and also be distinct from the @{text "preds"}.
183
+ − 103
184
+ − 104
The first function constructs the term for one particular predicate, say
+ − 105
@{text "pred"}; the number of arguments of this predicate is
+ − 106
determined by the number of argument types of @{text "arg_tys"}.
+ − 107
So it takes these two parameters as arguments. The other arguments are
+ − 108
all the @{text "preds"} and the @{text "orules"}.
164
+ − 109
*}
+ − 110
176
+ − 111
ML %linenosgray{*fun defs_aux lthy orules preds (pred, arg_tys) =
164
+ − 112
let
+ − 113
fun mk_all x P = HOLogic.all_const (fastype_of x) $ lambda x P
+ − 114
+ − 115
val fresh_args =
176
+ − 116
arg_tys
164
+ − 117
|> map (pair "z")
179
+ − 118
|> Variable.variant_frees lthy (preds @ orules)
164
+ − 119
|> map Free
+ − 120
in
+ − 121
list_comb (pred, fresh_args)
+ − 122
|> fold_rev (curry HOLogic.mk_imp) orules
+ − 123
|> fold_rev mk_all preds
+ − 124
|> fold_rev lambda fresh_args
+ − 125
end*}
+ − 126
+ − 127
text {*
184
+ − 128
The function in Line 3 is just a helper function for constructing universal
+ − 129
quantifications. The code in Lines 5 to 9 produces the fresh @{text
+ − 130
"\<^raw:$zs$>"}. For this it pairs every argument type with the string
+ − 131
@{text [quotes] "z"} (Line 7); then generates variants for all these strings
+ − 132
so that they are unique w.r.t.~to the @{text "orules"} and the predicates;
+ − 133
in Line 9 it generates the corresponding variable terms for the unique
+ − 134
strings.
164
+ − 135
183
+ − 136
The unique free variables are applied to the predicate (Line 11) using the
+ − 137
function @{ML list_comb}; then the @{text orules} are prefixed (Line 12); in
+ − 138
Line 13 we quantify over all predicates; and in line 14 we just abstract
+ − 139
over all the @{text "\<^raw:$zs$>"}, i.e.~the fresh arguments of the
+ − 140
predicate.
179
+ − 141
+ − 142
A testcase for this function is
164
+ − 143
*}
+ − 144
179
+ − 145
local_setup %gray{* fn lthy =>
176
+ − 146
let
+ − 147
val orules = [@{prop "even 0"},
+ − 148
@{prop "\<forall>n::nat. odd n \<longrightarrow> even (Suc n)"},
+ − 149
@{prop "\<forall>n::nat. even n \<longrightarrow> odd (Suc n)"}]
184
+ − 150
val preds = [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}, @{term "z::nat"}]
176
+ − 151
val pred = @{term "even::nat\<Rightarrow>bool"}
+ − 152
val arg_tys = [@{typ "nat"}]
+ − 153
val def = defs_aux lthy orules preds (pred, arg_tys)
173
+ − 154
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
+ − 155
warning (Syntax.string_of_term lthy def); lthy
179
+ − 156
end *}
173
+ − 157
91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 158
text {*
184
+ − 159
It constructs the left-hand side for the definition of @{text "even"}. So we obtain
183
+ − 160
as printout the term
179
+ − 161
+ − 162
@{text [display]
+ − 163
"\<lambda>z. \<forall>even odd. (even 0) \<longrightarrow> (\<forall>n. odd n \<longrightarrow> even (Suc n))
+ − 164
\<longrightarrow> (\<forall>n. even n \<longrightarrow> odd (Suc n)) \<longrightarrow> even z"}
+ − 165
184
+ − 166
The main function for the definitions now has to just iterate the function
+ − 167
@{ML defs_aux} over all predicates. The argument @{text "preds"} is again
+ − 168
the the list of predicates as @{ML_type term}s; the argument @{text
+ − 169
"prednames"} is the list of names of the predicates; @{text "arg_tyss"} is
+ − 170
the list of argument-type-lists for each predicate.
164
+ − 171
*}
+ − 172
+ − 173
ML %linenosgray{*fun definitions rules preds prednames syns arg_typss lthy =
+ − 174
let
+ − 175
val thy = ProofContext.theory_of lthy
+ − 176
val orules = map (ObjectLogic.atomize_term thy) rules
+ − 177
val defs = map (defs_aux lthy orules preds) (preds ~~ arg_typss)
+ − 178
in
+ − 179
fold_map make_defs (prednames ~~ syns ~~ defs) lthy
+ − 180
end*}
+ − 181
+ − 182
text {*
184
+ − 183
The user will state the introduction rules using meta-implications and
+ − 184
meta-quanti\-fications. In Line 4, we transform these introduction rules into
+ − 185
the object logic (since definitions cannot be stated with
+ − 186
meta-connectives). To do this transformation we have to obtain the theory
+ − 187
behind the local theory (Line 3); with this theory we can use the function
+ − 188
@{ML ObjectLogic.atomize_term} to make the transformation (Line 4). The call
+ − 189
to @{ML defs_aux} in Line 5 produces all left-hand sides of the
+ − 190
definitions. The actual definitions are then made in Line 7. The result
+ − 191
of the function is a list of theorems and a local theory.
+ − 192
179
+ − 193
+ − 194
A testcase for this function is
164
+ − 195
*}
+ − 196
179
+ − 197
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
+ − 198
let
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 199
val rules = [@{prop "even 0"},
179
+ − 200
@{prop "\<And>n::nat. odd n \<Longrightarrow> even (Suc n)"},
+ − 201
@{prop "\<And>n::nat. even n \<Longrightarrow> odd (Suc n)"}]
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 202
val preds = [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}]
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 203
val prednames = [@{binding "even"}, @{binding "odd"}]
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 204
val syns = [NoSyn, NoSyn]
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 205
val arg_tyss = [[@{typ "nat"}], [@{typ "nat"}]]
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 206
val (defs, lthy') = definitions rules preds prednames syns arg_tyss lthy
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 207
in
183
+ − 208
warning (str_of_thms lthy' defs); lthy'
179
+ − 209
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
+ − 210
179
+ − 211
text {*
184
+ − 212
where we feed into the functions all parameters corresponding to
+ − 213
the @{text even}-@{text odd} example. The definitions we obtain
+ − 214
are:
179
+ − 215
183
+ − 216
\begin{isabelle}
+ − 217
\isacommand{thm}~@{text "even_def odd_def"}\\
+ − 218
@{text [break]
+ − 219
"> even \<equiv> \<lambda>z. \<forall>even odd. (even 0) \<longrightarrow> (\<forall>n. odd n \<longrightarrow> even (Suc n))
+ − 220
> \<longrightarrow> (\<forall>n. even n \<longrightarrow> odd (Suc n)) \<longrightarrow> even z,
+ − 221
> odd \<equiv> \<lambda>z. \<forall>even odd. (even 0) \<longrightarrow> (\<forall>n. odd n \<longrightarrow> even (Suc n))
184
+ − 222
> \<longrightarrow> (\<forall>n. even n \<longrightarrow> odd (Suc n)) \<longrightarrow> odd z"}
183
+ − 223
\end{isabelle}
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 224
164
+ − 225
184
+ − 226
This completes the code for making the definitions. Next we deal with
+ − 227
the induction principles. Recall that the proof of the induction principle
+ − 228
for @{text "even"} was:
179
+ − 229
*}
176
+ − 230
184
+ − 231
lemma man_ind_principle:
183
+ − 232
assumes prems: "even n"
+ − 233
shows "P 0 \<Longrightarrow> (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n"
176
+ − 234
apply(atomize (full))
+ − 235
apply(cut_tac prems)
+ − 236
apply(unfold even_def)
+ − 237
apply(drule spec[where x=P])
+ − 238
apply(drule spec[where x=Q])
+ − 239
apply(assumption)
+ − 240
done
+ − 241
179
+ − 242
text {*
184
+ − 243
The code for such induction principles has to accomplish two tasks:
+ − 244
constructing the induction principles from the given introduction
+ − 245
rules and then automatically generating a proof of them using a tactic.
183
+ − 246
+ − 247
The tactic will use the following helper function for instantiating universal
184
+ − 248
quantifiers.
179
+ − 249
*}
176
+ − 250
179
+ − 251
ML{*fun inst_spec ctrm =
+ − 252
Drule.instantiate' [SOME (ctyp_of_term ctrm)] [NONE, SOME ctrm] @{thm spec}*}
164
+ − 253
+ − 254
text {*
184
+ − 255
This helper function instantiates the @{text "?x"} in the theorem
+ − 256
@{thm spec} with a given @{ML_type cterm}. Together with the tactic
164
+ − 257
*}
+ − 258
183
+ − 259
ML{*fun inst_spec_tac ctrms =
+ − 260
EVERY' (map (dtac o inst_spec) ctrms)*}
+ − 261
+ − 262
text {*
184
+ − 263
we can use @{ML inst_spec} in the following proof to instantiate the
+ − 264
three quantifiers in the assumption.
+ − 265
*}
183
+ − 266
184
+ − 267
lemma
+ − 268
fixes P::"nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
+ − 269
shows "\<forall>x y z. P x y z \<Longrightarrow> True"
176
+ − 270
apply (tactic {*
184
+ − 271
inst_spec_tac [@{cterm "a::nat"},@{cterm "b::nat"},@{cterm "c::nat"}] 1 *})
179
+ − 272
txt {*
183
+ − 273
We obtain the goal state
179
+ − 274
+ − 275
\begin{minipage}{\textwidth}
+ − 276
@{subgoals}
+ − 277
\end{minipage}*}
164
+ − 278
(*<*)oops(*>*)
+ − 279
+ − 280
text {*
183
+ − 281
Now the complete tactic for proving the induction principles can
+ − 282
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
+ − 283
*}
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 284
180
+ − 285
ML %linenosgray{*fun induction_tac defs prems insts =
176
+ − 286
EVERY1 [ObjectLogic.full_atomize_tac,
164
+ − 287
cut_facts_tac prems,
+ − 288
K (rewrite_goals_tac defs),
183
+ − 289
inst_spec_tac insts,
164
+ − 290
assume_tac]*}
+ − 291
179
+ − 292
text {*
184
+ − 293
We only have to give it as arguments the definitions, the premise
+ − 294
(like @{text "even n"})
+ − 295
and the instantiations. Compare this with the manual proof given for the
+ − 296
lemma @{thm [source] man_ind_principle}.
+ − 297
A testcase for this tactic is the function
179
+ − 298
*}
+ − 299
183
+ − 300
ML{*fun test_tac prems =
+ − 301
let
+ − 302
val defs = [@{thm even_def}, @{thm odd_def}]
+ − 303
val insts = [@{cterm "P::nat\<Rightarrow>bool"}, @{cterm "Q::nat\<Rightarrow>bool"}]
+ − 304
in
+ − 305
induction_tac defs prems insts
+ − 306
end*}
+ − 307
+ − 308
text {*
184
+ − 309
which indeed proves the induction principle:
183
+ − 310
*}
+ − 311
164
+ − 312
lemma
183
+ − 313
assumes prems: "even n"
+ − 314
shows "P 0 \<Longrightarrow> (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n"
+ − 315
apply(tactic {* test_tac @{thms prems} *})
164
+ − 316
done
+ − 317
165
+ − 318
text {*
184
+ − 319
While the tactic for the induction principle is relatively simple,
+ − 320
it is a bit harder to construct the goals from the introduction
+ − 321
rules the user provides. In general we have to construct for each predicate
183
+ − 322
@{text "pred"} a goal of the form
179
+ − 323
+ − 324
@{text [display]
184
+ − 325
"\<And>\<^raw:$zs$>. pred \<^raw:$zs$> \<Longrightarrow> rules[preds := \<^raw:$Ps$>] \<Longrightarrow> \<^raw:$P$> \<^raw:$zs$>"}
180
+ − 326
183
+ − 327
where the given predicates @{text preds} are replaced in the introduction
184
+ − 328
rules by new distinct variables written @{text "\<^raw:$Ps$>"}.
+ − 329
We also need to generate fresh arguments for the predicate @{text "pred"} in
+ − 330
the premise and the @{text "\<^raw:$P$>"} in the conclusion. We achieve
+ − 331
that in two steps.
179
+ − 332
184
+ − 333
The function below expects that the introduction rules are already appropriately
+ − 334
substituted. The argument @{text "srules"} stands for these substituted
+ − 335
rules; @{text cnewpreds} are the certified terms coresponding
180
+ − 336
to the variables @{text "\<^raw:$Ps$>"}; @{text "pred"} is the predicate for
+ − 337
which we prove the introduction principle; @{text "newpred"} is its
183
+ − 338
replacement and @{text "tys"} are the argument types of this predicate.
165
+ − 339
*}
+ − 340
183
+ − 341
ML %linenosgray{*fun prove_induction lthy defs srules cnewpreds ((pred, newpred), 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
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 zs = replicate (length tys) "z"
179
+ − 344
val (newargnames, lthy') = Variable.variant_fixes zs 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
+ − 345
val newargs = map Free (newargnames ~~ tys)
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
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 352
(fn {prems, ...} => induction_tac defs prems cnewpreds)
179
+ − 353
|> singleton (ProofContext.export lthy' lthy)
+ − 354
end *}
+ − 355
180
+ − 356
text {*
184
+ − 357
In Line 3 we produce names @{text "\<^raw:$zs$>"} for each type in the
183
+ − 358
argument type list. Line 4 makes these names unique and declares them as
+ − 359
\emph{free} (but fixed) variables in the local theory @{text "lthy'"}. In
+ − 360
Line 5 we just construct the terms corresponding to these variables.
+ − 361
The term variables are applied to the predicate in Line 7 (this corresponds
+ − 362
to the first premise @{text "pred \<^raw:$zs$>"} of the induction principle).
+ − 363
In Line 8 and 9, we first construct the term @{text "\<^raw:$P$>\<^raw:$zs$>"}
184
+ − 364
and then add the (substituded) introduction rules as premises. In case that
+ − 365
no introduction rules are given, the conclusion of this implication needs
183
+ − 366
to be wrapped inside a @{term Trueprop}, otherwise the Isabelle's goal
+ − 367
mechanism will fail.
180
+ − 368
184
+ − 369
In Line 11 we set up the goal to be proved; in the next line call the tactic
+ − 370
for proving the induction principle. This tactic expects definitions, the
+ − 371
premise and the (certified) predicates with which the introduction rules
+ − 372
have been substituted. This will return a theorem. However, it is a theorem
+ − 373
proved inside the local theory @{text "lthy'"}, where the variables @{text
+ − 374
"\<^raw:$zs$>"} are fixed, but free. By exporting this theorem from @{text
+ − 375
"lthy'"} (which contains the @{text "\<^raw:$zs$>"} as free) to @{text
+ − 376
"lthy"} (which does not), we obtain the desired quantifications @{text
+ − 377
"\<And>\<^raw:$zs$>"}.
180
+ − 378
184
+ − 379
(FIXME testcase)
+ − 380
+ − 381
+ − 382
Now it is left to produce the new predicates with which the introduction
183
+ − 383
rules are substituted.
180
+ − 384
*}
165
+ − 385
184
+ − 386
ML %linenosgray{*fun inductions rules defs preds arg_tyss lthy =
164
+ − 387
let
+ − 388
val Ps = replicate (length preds) "P"
183
+ − 389
val (newprednames, lthy') = Variable.variant_fixes Ps lthy
164
+ − 390
183
+ − 391
val thy = ProofContext.theory_of lthy'
164
+ − 392
184
+ − 393
val tyss' = map (fn tys => tys ---> HOLogic.boolT) arg_tyss
165
+ − 394
val newpreds = map Free (newprednames ~~ tyss')
164
+ − 395
val cnewpreds = map (cterm_of thy) newpreds
184
+ − 396
val srules = map (subst_free (preds ~~ newpreds)) rules
164
+ − 397
+ − 398
in
184
+ − 399
map (prove_induction lthy' defs srules cnewpreds)
+ − 400
(preds ~~ newpreds ~~ arg_tyss)
183
+ − 401
|> ProofContext.export lthy' lthy
165
+ − 402
end*}
+ − 403
184
+ − 404
text {*
+ − 405
In Line 3 we generate a string @{text [quotes] "P"} for each predicate.
+ − 406
In Line 4, we use the same trick as in the previous function, that is making the
+ − 407
@{text "\<^raw:$Ps$>"} fresh and declaring them as fixed but free in
+ − 408
the new local theory @{text "lthy'"}. From the local theory we extract
+ − 409
the ambient theory in Line 6. We need this theory in order to certify
+ − 410
the new predicates. In Line 8 we calculate the types of these new predicates
+ − 411
using the argument types. Next we turn them into terms and subsequently
+ − 412
certify them. We can now produce the substituted introduction rules
+ − 413
(Line 11). Line 14 and 15 just iterate the proofs for all predicates.
+ − 414
From this we obtain a list of theorems. Finally we need to export the
+ − 415
fixed variables @{text "\<^raw:$Ps$>"} to obtain the correct quantification
+ − 416
(Line 16).
+ − 417
+ − 418
A testcase for this function is
+ − 419
*}
+ − 420
+ − 421
local_setup %gray {* fn lthy =>
+ − 422
let
+ − 423
val rules = [@{prop "even (0::nat)"},
+ − 424
@{prop "\<And>n::nat. odd n \<Longrightarrow> even (Suc n)"},
+ − 425
@{prop "\<And>n::nat. even n \<Longrightarrow> odd (Suc n)"}]
+ − 426
val defs = [@{thm even_def}, @{thm odd_def}]
+ − 427
val preds = [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}]
+ − 428
val tyss = [[@{typ "nat"}], [@{typ "nat"}]]
+ − 429
val ind_thms = inductions rules defs preds tyss lthy
+ − 430
in
+ − 431
warning (str_of_thms lthy ind_thms); lthy
+ − 432
end
165
+ − 433
*}
+ − 434
176
+ − 435
184
+ − 436
text {*
+ − 437
which prints out
+ − 438
+ − 439
@{text [display]
+ − 440
"> even z \<Longrightarrow>
+ − 441
> P 0 \<Longrightarrow> (\<And>m. Pa m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Pa (Suc m)) \<Longrightarrow> P z,
+ − 442
> odd z \<Longrightarrow>
+ − 443
> P 0 \<Longrightarrow> (\<And>m. Pa m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Pa (Suc m)) \<Longrightarrow> Pa z"}
+ − 444
+ − 445
+ − 446
This completes the code for the induction principles. Finally we can
+ − 447
prove the introduction rules.
+ − 448
+ − 449
*}
+ − 450
+ − 451
ML {* ObjectLogic.rulify *}
+ − 452
165
+ − 453
+ − 454
ML{*val all_elims = fold (fn ct => fn th => th RS inst_spec ct)
+ − 455
val imp_elims = fold (fn th => fn th' => [th', th] MRS @{thm mp})*}
+ − 456
+ − 457
ML{*fun subproof2 prem params2 prems2 =
+ − 458
SUBPROOF (fn {prems, ...} =>
+ − 459
let
+ − 460
val prem' = prems MRS prem;
+ − 461
val prem'' =
+ − 462
case prop_of prem' of
+ − 463
_ $ (Const (@{const_name All}, _) $ _) =>
+ − 464
prem' |> all_elims params2
+ − 465
|> imp_elims prems2
+ − 466
| _ => prem';
+ − 467
in
+ − 468
rtac prem'' 1
+ − 469
end)*}
+ − 470
+ − 471
ML{*fun subproof1 rules preds i =
+ − 472
SUBPROOF (fn {params, prems, context = ctxt', ...} =>
+ − 473
let
+ − 474
val (prems1, prems2) = chop (length prems - length rules) prems;
+ − 475
val (params1, params2) = chop (length params - length preds) params;
+ − 476
in
+ − 477
rtac (ObjectLogic.rulify (all_elims params1 (nth prems2 i))) 1
+ − 478
THEN
+ − 479
EVERY1 (map (fn prem => subproof2 prem params2 prems2 ctxt') prems1)
+ − 480
end)*}
+ − 481
+ − 482
ML{*
+ − 483
fun introductions_tac defs rules preds i ctxt =
+ − 484
EVERY1 [ObjectLogic.rulify_tac,
+ − 485
K (rewrite_goals_tac defs),
184
+ − 486
REPEAT o (resolve_tac [@{thm allI}, @{thm impI}]),
165
+ − 487
subproof1 rules preds i ctxt]*}
+ − 488
173
+ − 489
lemma evenS:
+ − 490
shows "odd m \<Longrightarrow> even (Suc m)"
+ − 491
apply(tactic {*
+ − 492
let
+ − 493
val rules = [@{prop "even (0::nat)"},
+ − 494
@{prop "\<And>n::nat. odd n \<Longrightarrow> even (Suc n)"},
+ − 495
@{prop "\<And>n::nat. even n \<Longrightarrow> odd (Suc n)"}]
+ − 496
val defs = [@{thm even_def}, @{thm odd_def}]
+ − 497
val preds = [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}]
+ − 498
in
+ − 499
introductions_tac defs rules preds 1 @{context}
+ − 500
end *})
+ − 501
done
+ − 502
165
+ − 503
ML{*fun introductions rules preds defs lthy =
+ − 504
let
+ − 505
fun prove_intro (i, goal) =
+ − 506
Goal.prove lthy [] [] goal
+ − 507
(fn {context, ...} => introductions_tac defs rules preds i context)
+ − 508
in
+ − 509
map_index prove_intro rules
164
+ − 510
end*}
+ − 511
176
+ − 512
text {* main internal function *}
+ − 513
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 514
ML %linenosgray{*fun add_inductive pred_specs rule_specs lthy =
165
+ − 515
let
+ − 516
val syns = map snd pred_specs
+ − 517
val pred_specs' = map fst pred_specs
+ − 518
val prednames = map fst pred_specs'
+ − 519
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
+ − 520
165
+ − 521
val tyss = map (binder_types o fastype_of) preds
+ − 522
val (attrs, rules) = split_list rule_specs
+ − 523
+ − 524
val (defs, lthy') = definitions rules preds prednames syns tyss lthy
+ − 525
val ind_rules = inductions rules defs preds tyss lthy'
+ − 526
val intro_rules = introductions rules preds defs lthy'
91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 527
165
+ − 528
val mut_name = space_implode "_" (map Binding.name_of prednames)
+ − 529
val case_names = map (Binding.name_of o fst) attrs
+ − 530
in
+ − 531
lthy'
+ − 532
|> LocalTheory.notes Thm.theoremK (map (fn (((a, atts), _), th) =>
+ − 533
((Binding.qualify false mut_name a, atts), [([th], [])])) (rule_specs ~~ intro_rules))
+ − 534
|-> (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
+ − 535
((Binding.qualify false mut_name (@{binding "intros"}), []), maps snd intross))
165
+ − 536
|>> snd
+ − 537
||>> (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
+ − 538
((Binding.qualify false (Binding.name_of R) (@{binding "induct"}),
165
+ − 539
[Attrib.internal (K (RuleCases.case_names case_names)),
+ − 540
Attrib.internal (K (RuleCases.consumes 1)),
+ − 541
Attrib.internal (K (Induct.induct_pred ""))]), [([th], [])]))
+ − 542
(pred_specs ~~ ind_rules)) #>> maps snd)
+ − 543
|> snd
+ − 544
end*}
91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 545
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 546
ML{*fun add_inductive_cmd pred_specs rule_specs lthy =
165
+ − 547
let
183
+ − 548
val ((pred_specs', rule_specs'), _) =
+ − 549
Specification.read_spec pred_specs rule_specs lthy
165
+ − 550
in
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 551
add_inductive pred_specs' rule_specs' lthy
165
+ − 552
end*}
+ − 553
+ − 554
ML{*val spec_parser =
+ − 555
OuterParse.fixes --
+ − 556
Scan.optional
+ − 557
(OuterParse.$$$ "where" |--
+ − 558
OuterParse.!!!
+ − 559
(OuterParse.enum1 "|"
+ − 560
(SpecParse.opt_thm_name ":" -- OuterParse.prop))) []*}
+ − 561
+ − 562
ML{*val specification =
+ − 563
spec_parser >>
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 564
(fn ((pred_specs), rule_specs) => add_inductive_cmd pred_specs rule_specs)*}
165
+ − 565
185
+ − 566
ML{*val _ = OuterSyntax.local_theory "simple_inductive"
+ − 567
"define inductive predicates"
+ − 568
OuterKeyword.thy_decl specification*}
91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 569
124
+ − 570
text {*
+ − 571
Things to include at the end:
+ − 572
+ − 573
\begin{itemize}
+ − 574
\item say something about add-inductive-i to return
+ − 575
the rules
+ − 576
\item say that the induction principle is weaker (weaker than
+ − 577
what the standard inductive package generates)
+ − 578
\end{itemize}
+ − 579
+ − 580
*}
+ − 581
165
+ − 582
simple_inductive
+ − 583
Even and Odd
+ − 584
where
+ − 585
Even0: "Even 0"
+ − 586
| EvenS: "Odd n \<Longrightarrow> Even (Suc n)"
+ − 587
| OddS: "Even n \<Longrightarrow> Odd (Suc n)"
124
+ − 588
91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 589
end