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
subsection {* Definitions *}
+ − 8
+ − 9
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
+ − 10
@{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
+ − 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] "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
+ − 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] "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
+ − 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] "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
+ − 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] "oind ::= \<forall>zs. pred zs \<longrightarrow> orules[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
+ − 19
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 20
So we have @{text "pred zs"} and @{text "orules[preds::=Ps]"}; have to show
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 21
@{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
+ − 22
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
+ − 23
@{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
+ − 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
We have to show @{text "\<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
+ − 26
expanding the defs
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 27
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 28
@{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
+ − 29
"\<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
+ − 30
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 31
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
+ − 32
@{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
+ − 33
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 34
the @{text "orules"} are of the form @{text "\<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
+ − 35
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 36
using the @{text "As"} we ????
164
+ − 37
*}
+ − 38
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 39
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 40
text {*
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 41
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
+ − 42
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 43
@{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
+ − 44
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 45
We use the following wrapper function to actually make the definition via
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 46
@{ML LocalTheory.define}. The function takes a predicate name, a syntax
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 47
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
+ − 48
*}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 49
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 50
ML %linenosgray{*fun make_defs ((predname, syn), trm) lthy =
164
+ − 51
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
+ − 52
val arg = ((predname, syn), (Attrib.empty_binding, trm))
176
+ − 53
val ((_, (_ , thm)), lthy') = LocalTheory.define Thm.internalK arg lthy
164
+ − 54
in
176
+ − 55
(thm, lthy')
164
+ − 56
end*}
+ − 57
+ − 58
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
+ − 59
Returns the definition (as theorem) and the local theory in which this definition has
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 60
been made. The @{ML internalK in Thm} in Line 4 is just a flag attached to the
176
+ − 61
theorem (others flags are @{ML definitionK in Thm} or @{ML axiomK in Thm}).
+ − 62
These flags just classify theorems and have no significant meaning, except
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 63
for tools such as finding theorems in the theorem database. We also
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 64
use @{ML empty_binding in Attrib} in Line 3, since the definition does not need any
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 65
theorem attributes.
164
+ − 66
*}
+ − 67
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 68
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
+ − 69
let
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 70
val arg = ((Binding.name "MyTrue", NoSyn), @{term True})
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 71
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
+ − 72
val _ = warning (str_of_thm lthy' def)
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 73
in
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 74
lthy'
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 75
end *}
164
+ − 76
+ − 77
text {*
176
+ − 78
Prints out the theorem @{prop "MyTrue \<equiv> True"}.
164
+ − 79
*}
+ − 80
+ − 81
text {*
+ − 82
Constructs the term for the definition: the main arguments are a predicate
+ − 83
and the types of the arguments, it also expects orules which are the
+ − 84
intro rules in the object logic; preds which are all predicates. returns the
+ − 85
term.
+ − 86
*}
+ − 87
176
+ − 88
ML %linenosgray{*fun defs_aux lthy orules preds (pred, arg_tys) =
164
+ − 89
let
+ − 90
fun mk_all x P = HOLogic.all_const (fastype_of x) $ lambda x P
+ − 91
+ − 92
val fresh_args =
176
+ − 93
arg_tys
164
+ − 94
|> map (pair "z")
+ − 95
|> Variable.variant_frees lthy orules
+ − 96
|> map Free
+ − 97
in
+ − 98
list_comb (pred, fresh_args)
+ − 99
|> fold_rev (curry HOLogic.mk_imp) orules
+ − 100
|> fold_rev mk_all preds
+ − 101
|> fold_rev lambda fresh_args
+ − 102
end*}
+ − 103
+ − 104
text {*
+ − 105
The lines 5-9 produce fresh arguments with which the predicate can be applied.
+ − 106
For this it pairs every type with a string @{text [quotes] "z"} (Line 7); then
+ − 107
generates variants for all these strings (names) so that they are unique w.r.t.~to
+ − 108
the intro rules; in Line 9 it generates the corresponding variable terms for these
+ − 109
unique names.
+ − 110
+ − 111
The unique free variables are applied to the predicate (Line 11); then
+ − 112
the intro rules are attached as preconditions (Line 12); in Line 13 we
+ − 113
quantify over all predicates; and in line 14 we just abstract over all
+ − 114
the (fresh) arguments of the predicate.
+ − 115
*}
+ − 116
176
+ − 117
+ − 118
local_setup {*
+ − 119
fn lthy =>
+ − 120
let
+ − 121
val orules = [@{prop "even 0"},
+ − 122
@{prop "\<forall>n::nat. odd n \<longrightarrow> even (Suc n)"},
+ − 123
@{prop "\<forall>n::nat. even n \<longrightarrow> odd (Suc n)"}]
173
+ − 124
val preds = [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}]
176
+ − 125
val pred = @{term "even::nat\<Rightarrow>bool"}
+ − 126
val arg_tys = [@{typ "nat"}]
+ − 127
val def = defs_aux lthy orules preds (pred, arg_tys)
173
+ − 128
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
+ − 129
warning (Syntax.string_of_term lthy def); lthy
173
+ − 130
end*}
+ − 131
91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 132
text {*
164
+ − 133
The arguments of the main function for the definitions are
+ − 134
the intro rules; the predicates and their names, their syntax
+ − 135
annotations and the argument types of each predicate. It
+ − 136
returns a theorem list (the definitions) and a local
+ − 137
theory where the definitions are made
+ − 138
*}
+ − 139
+ − 140
ML %linenosgray{*fun definitions rules preds prednames syns arg_typss lthy =
+ − 141
let
+ − 142
val thy = ProofContext.theory_of lthy
+ − 143
val orules = map (ObjectLogic.atomize_term thy) rules
+ − 144
val defs = map (defs_aux lthy orules preds) (preds ~~ arg_typss)
+ − 145
in
+ − 146
fold_map make_defs (prednames ~~ syns ~~ defs) lthy
+ − 147
end*}
+ − 148
+ − 149
text {*
+ − 150
In line 4 we generate the intro rules in the object logic; for this we have to
+ − 151
obtain the theory behind the local theory (Line 3); with this we can
+ − 152
call @{ML defs_aux} to generate the terms for the left-hand sides.
+ − 153
The actual definitions are made in Line 7.
+ − 154
*}
+ − 155
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 156
local_setup {*
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 157
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
+ − 158
let
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 159
val rules = [@{prop "even 0"},
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 160
@{prop "\<And>n::nat. odd n \<Longrightarrow> even (Suc n)"},
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 161
@{prop "\<And>n::nat. even n \<Longrightarrow> odd (Suc n)"}]
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 162
val preds = [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}]
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 163
val prednames = [Binding.name "even", Binding.name "odd"]
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 164
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
+ − 165
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
+ − 166
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
+ − 167
in
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 168
warning (str_of_thms lthy' defs); lthy
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 169
end*}
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 170
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 171
165
+ − 172
subsection {* Induction Principles *}
164
+ − 173
176
+ − 174
text {* Recall the proof for the induction principle for @{term "even"}: *}
+ − 175
+ − 176
lemma
+ − 177
assumes prems: "even n"
+ − 178
shows "P 0 \<Longrightarrow>
+ − 179
(\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n"
+ − 180
apply(atomize (full))
+ − 181
apply(cut_tac prems)
+ − 182
apply(unfold even_def)
+ − 183
apply(drule spec[where x=P])
+ − 184
apply(drule spec[where x=Q])
+ − 185
apply(assumption)
+ − 186
done
+ − 187
+ − 188
+ − 189
text {* We need to be able to instantiate universal quantifiers. *}
+ − 190
164
+ − 191
ML{*fun inst_spec ct =
+ − 192
Drule.instantiate' [SOME (ctyp_of_term ct)] [NONE, SOME ct] @{thm spec}*}
+ − 193
+ − 194
text {*
165
+ − 195
Instantiates the @{text "?x"} in @{thm spec} with a @{ML_type cterm}.
164
+ − 196
*}
+ − 197
+ − 198
text {*
+ − 199
Instantiates universal qantifications in the premises
+ − 200
*}
+ − 201
+ − 202
lemma "\<forall>x\<^isub>1 x\<^isub>2 x\<^isub>3. P (x\<^isub>1::nat) (x\<^isub>2::nat) (x\<^isub>3::nat) \<Longrightarrow> True"
176
+ − 203
apply (tactic {*
+ − 204
let
+ − 205
val ctrms = [@{cterm "y\<^isub>1::nat"},@{cterm "y\<^isub>2::nat"},@{cterm "y\<^isub>3::nat"}]
+ − 206
in
+ − 207
EVERY1 (map (dtac o inst_spec) ctrms)
+ − 208
end *})
165
+ − 209
txt {* \begin{minipage}{\textwidth}
+ − 210
@{subgoals}
+ − 211
\end{minipage}*}
164
+ − 212
(*<*)oops(*>*)
+ − 213
+ − 214
text {*
176
+ − 215
Now the tactic for proving the induction rules:
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
+ − 216
*}
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 217
164
+ − 218
ML{*fun induction_tac defs prems insts =
176
+ − 219
EVERY1 [ObjectLogic.full_atomize_tac,
164
+ − 220
cut_facts_tac prems,
+ − 221
K (rewrite_goals_tac defs),
+ − 222
EVERY' (map (dtac o inst_spec) insts),
+ − 223
assume_tac]*}
+ − 224
+ − 225
lemma
176
+ − 226
assumes prems: "even n"
164
+ − 227
shows "P 0 \<Longrightarrow>
+ − 228
(\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n"
165
+ − 229
apply(tactic {*
+ − 230
let
+ − 231
val defs = [@{thm even_def}, @{thm odd_def}]
+ − 232
val insts = [@{cterm "P::nat\<Rightarrow>bool"}, @{cterm "Q::nat\<Rightarrow>bool"}]
+ − 233
in
+ − 234
induction_tac defs @{thms prems} insts
+ − 235
end *})
164
+ − 236
done
+ − 237
176
+ − 238
165
+ − 239
text {*
+ − 240
While the generic proof is relatively simple, it is a bit harder to
+ − 241
set up the goals for the induction principles.
+ − 242
*}
+ − 243
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 244
ML {*
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 245
fun prove_induction lthy2 defs rules cnewpreds ((pred, newpred), tys) =
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 246
let
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 247
val zs = replicate (length tys) "z"
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 248
val (newargnames, lthy3) = Variable.variant_fixes zs lthy2;
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 249
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
+ − 250
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 251
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
+ − 252
val goal = Logic.list_implies
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 253
(rules, HOLogic.mk_Trueprop (list_comb (newpred, newargs)))
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 254
in
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 255
Goal.prove lthy3 [] [prem] goal
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 256
(fn {prems, ...} => induction_tac defs prems cnewpreds)
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 257
|> singleton (ProofContext.export lthy3 lthy2)
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 258
end
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 259
*}
165
+ − 260
+ − 261
ML %linenosgray{*fun inductions rules defs preds tyss lthy1 =
164
+ − 262
let
+ − 263
val Ps = replicate (length preds) "P"
+ − 264
val (newprednames, lthy2) = Variable.variant_fixes Ps lthy1
+ − 265
+ − 266
val thy = ProofContext.theory_of lthy2
+ − 267
165
+ − 268
val tyss' = map (fn tys => tys ---> HOLogic.boolT) tyss
+ − 269
val newpreds = map Free (newprednames ~~ tyss')
164
+ − 270
val cnewpreds = map (cterm_of thy) newpreds
+ − 271
val rules' = map (subst_free (preds ~~ newpreds)) rules
+ − 272
+ − 273
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
+ − 274
map (prove_induction lthy2 defs rules' cnewpreds)
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 275
(preds ~~ newpreds ~~ tyss)
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 276
|> ProofContext.export lthy2 lthy1
165
+ − 277
end*}
+ − 278
+ − 279
ML {*
+ − 280
let
173
+ − 281
val rules = [@{prop "even (0::nat)"},
+ − 282
@{prop "\<And>n::nat. odd n \<Longrightarrow> even (Suc n)"},
+ − 283
@{prop "\<And>n::nat. even n \<Longrightarrow> odd (Suc n)"}]
165
+ − 284
val defs = [@{thm even_def}, @{thm odd_def}]
+ − 285
val preds = [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}]
176
+ − 286
val tyss = [[@{typ "nat"}], [@{typ "nat"}]]
165
+ − 287
in
+ − 288
inductions rules defs preds tyss @{context}
+ − 289
end
+ − 290
*}
+ − 291
176
+ − 292
165
+ − 293
subsection {* Introduction Rules *}
+ − 294
+ − 295
ML{*val all_elims = fold (fn ct => fn th => th RS inst_spec ct)
+ − 296
val imp_elims = fold (fn th => fn th' => [th', th] MRS @{thm mp})*}
+ − 297
+ − 298
ML{*fun subproof2 prem params2 prems2 =
+ − 299
SUBPROOF (fn {prems, ...} =>
+ − 300
let
+ − 301
val prem' = prems MRS prem;
+ − 302
val prem'' =
+ − 303
case prop_of prem' of
+ − 304
_ $ (Const (@{const_name All}, _) $ _) =>
+ − 305
prem' |> all_elims params2
+ − 306
|> imp_elims prems2
+ − 307
| _ => prem';
+ − 308
in
+ − 309
rtac prem'' 1
+ − 310
end)*}
+ − 311
+ − 312
ML{*fun subproof1 rules preds i =
+ − 313
SUBPROOF (fn {params, prems, context = ctxt', ...} =>
+ − 314
let
+ − 315
val (prems1, prems2) = chop (length prems - length rules) prems;
+ − 316
val (params1, params2) = chop (length params - length preds) params;
+ − 317
in
+ − 318
rtac (ObjectLogic.rulify (all_elims params1 (nth prems2 i))) 1
+ − 319
THEN
+ − 320
EVERY1 (map (fn prem => subproof2 prem params2 prems2 ctxt') prems1)
+ − 321
end)*}
+ − 322
+ − 323
ML{*
+ − 324
fun introductions_tac defs rules preds i ctxt =
+ − 325
EVERY1 [ObjectLogic.rulify_tac,
+ − 326
K (rewrite_goals_tac defs),
+ − 327
REPEAT o (resolve_tac [@{thm allI},@{thm impI}]),
+ − 328
subproof1 rules preds i ctxt]*}
+ − 329
173
+ − 330
lemma evenS:
+ − 331
shows "odd m \<Longrightarrow> even (Suc m)"
+ − 332
apply(tactic {*
+ − 333
let
+ − 334
val rules = [@{prop "even (0::nat)"},
+ − 335
@{prop "\<And>n::nat. odd n \<Longrightarrow> even (Suc n)"},
+ − 336
@{prop "\<And>n::nat. even n \<Longrightarrow> odd (Suc n)"}]
+ − 337
val defs = [@{thm even_def}, @{thm odd_def}]
+ − 338
val preds = [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}]
+ − 339
in
+ − 340
introductions_tac defs rules preds 1 @{context}
+ − 341
end *})
+ − 342
done
+ − 343
165
+ − 344
ML{*fun introductions rules preds defs lthy =
+ − 345
let
+ − 346
fun prove_intro (i, goal) =
+ − 347
Goal.prove lthy [] [] goal
+ − 348
(fn {context, ...} => introductions_tac defs rules preds i context)
+ − 349
in
+ − 350
map_index prove_intro rules
164
+ − 351
end*}
+ − 352
176
+ − 353
text {* main internal function *}
+ − 354
165
+ − 355
ML %linenosgray{*fun add_inductive_i pred_specs rule_specs lthy =
+ − 356
let
+ − 357
val syns = map snd pred_specs
+ − 358
val pred_specs' = map fst pred_specs
+ − 359
val prednames = map fst pred_specs'
+ − 360
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
+ − 361
165
+ − 362
val tyss = map (binder_types o fastype_of) preds
+ − 363
val (attrs, rules) = split_list rule_specs
+ − 364
+ − 365
val (defs, lthy') = definitions rules preds prednames syns tyss lthy
+ − 366
val ind_rules = inductions rules defs preds tyss lthy'
+ − 367
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
+ − 368
165
+ − 369
val mut_name = space_implode "_" (map Binding.name_of prednames)
+ − 370
val case_names = map (Binding.name_of o fst) attrs
+ − 371
in
+ − 372
lthy'
+ − 373
|> LocalTheory.notes Thm.theoremK (map (fn (((a, atts), _), th) =>
+ − 374
((Binding.qualify false mut_name a, atts), [([th], [])])) (rule_specs ~~ intro_rules))
+ − 375
|-> (fn intross => LocalTheory.note Thm.theoremK
+ − 376
((Binding.qualify false mut_name (Binding.name "intros"), []), maps snd intross))
+ − 377
|>> snd
+ − 378
||>> (LocalTheory.notes Thm.theoremK (map (fn (((R, _), _), th) =>
+ − 379
((Binding.qualify false (Binding.name_of R) (Binding.name "induct"),
+ − 380
[Attrib.internal (K (RuleCases.case_names case_names)),
+ − 381
Attrib.internal (K (RuleCases.consumes 1)),
+ − 382
Attrib.internal (K (Induct.induct_pred ""))]), [([th], [])]))
+ − 383
(pred_specs ~~ ind_rules)) #>> maps snd)
+ − 384
|> snd
+ − 385
end*}
91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 386
165
+ − 387
ML{*fun read_specification' vars specs lthy =
+ − 388
let
176
+ − 389
val specs' = map (fn (a, s) => (a, [s])) specs
165
+ − 390
val ((varst, specst), _) =
+ − 391
Specification.read_specification vars specs' lthy
+ − 392
val specst' = map (apsnd the_single) specst
+ − 393
in
+ − 394
(varst, specst')
+ − 395
end*}
+ − 396
+ − 397
ML{*fun add_inductive pred_specs rule_specs lthy =
+ − 398
let
+ − 399
val (pred_specs', rule_specs') =
+ − 400
read_specification' pred_specs rule_specs lthy
+ − 401
in
+ − 402
add_inductive_i pred_specs' rule_specs' lthy
+ − 403
end*}
+ − 404
+ − 405
ML{*val spec_parser =
+ − 406
OuterParse.opt_target --
+ − 407
OuterParse.fixes --
+ − 408
Scan.optional
+ − 409
(OuterParse.$$$ "where" |--
+ − 410
OuterParse.!!!
+ − 411
(OuterParse.enum1 "|"
+ − 412
(SpecParse.opt_thm_name ":" -- OuterParse.prop))) []*}
+ − 413
+ − 414
ML{*val specification =
+ − 415
spec_parser >>
+ − 416
(fn ((loc, pred_specs), rule_specs) =>
+ − 417
Toplevel.local_theory loc (add_inductive pred_specs rule_specs))*}
+ − 418
+ − 419
ML{*val _ = OuterSyntax.command "simple_inductive" "define inductive predicates"
+ − 420
OuterKeyword.thy_decl specification*}
91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 421
124
+ − 422
text {*
+ − 423
Things to include at the end:
+ − 424
+ − 425
\begin{itemize}
+ − 426
\item say something about add-inductive-i to return
+ − 427
the rules
+ − 428
\item say that the induction principle is weaker (weaker than
+ − 429
what the standard inductive package generates)
+ − 430
\end{itemize}
+ − 431
+ − 432
*}
+ − 433
165
+ − 434
simple_inductive
+ − 435
Even and Odd
+ − 436
where
+ − 437
Even0: "Even 0"
+ − 438
| EvenS: "Odd n \<Longrightarrow> Even (Suc n)"
+ − 439
| OddS: "Even n \<Longrightarrow> Odd (Suc n)"
124
+ − 440
91
667a0943c40b
added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 441
end