15
9da9ba2b095b
added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 1
theory Solutions
172
+ − 2
imports Base "Recipes/Timing"
15
9da9ba2b095b
added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 3
begin
9da9ba2b095b
added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 4
156
+ − 5
chapter {* Solutions to Most Exercises\label{ch:solutions} *}
15
9da9ba2b095b
added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 6
9da9ba2b095b
added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 7
text {* \solution{fun:revsum} *}
9da9ba2b095b
added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 8
293
+ − 9
ML{*fun rev_sum ((p as Const (@{const_name plus}, _)) $ t $ u) =
+ − 10
p $ u $ rev_sum t
271
+ − 11
| rev_sum t = t *}
+ − 12
293
+ − 13
text {*
316
+ − 14
An alternative solution using the function @{ML_ind mk_binop in HOLogic} is:
293
+ − 15
*}
271
+ − 16
69
+ − 17
ML{*fun rev_sum t =
47
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 18
let
130
+ − 19
fun dest_sum (Const (@{const_name plus}, _) $ u $ u') = u' :: dest_sum u
+ − 20
| dest_sum u = [u]
+ − 21
in
47
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 22
foldl1 (HOLogic.mk_binop @{const_name plus}) (dest_sum t)
130
+ − 23
end *}
15
9da9ba2b095b
added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 24
9da9ba2b095b
added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 25
text {* \solution{fun:makesum} *}
9da9ba2b095b
added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 26
69
+ − 27
ML{*fun make_sum t1 t2 =
130
+ − 28
HOLogic.mk_nat (HOLogic.dest_nat t1 + HOLogic.dest_nat t2) *}
15
9da9ba2b095b
added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 29
313
+ − 30
text {* \solution{ex:debruijn} *}
+ − 31
+ − 32
ML{*fun P n = @{term "P::nat \<Rightarrow> bool"} $ (HOLogic.mk_number @{typ "nat"} n)
+ − 33
+ − 34
fun rhs 1 = P 1
+ − 35
| rhs n = HOLogic.mk_conj (P n, rhs (n - 1))
+ − 36
+ − 37
fun lhs 1 n = HOLogic.mk_imp (HOLogic.mk_eq (P 1, P n), rhs n)
+ − 38
| lhs m n = HOLogic.mk_conj (HOLogic.mk_imp
+ − 39
(HOLogic.mk_eq (P (m - 1), P m), rhs n), lhs (m - 1) n)
+ − 40
+ − 41
fun de_bruijn n =
+ − 42
HOLogic.mk_Trueprop (HOLogic.mk_imp (lhs n n, rhs n))*}
+ − 43
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 44
text {* \solution{ex:scancmts} *}
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 45
130
+ − 46
ML{*val any = Scan.one (Symbol.not_eof)
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 47
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 48
val scan_cmt =
168
+ − 49
let
+ − 50
val begin_cmt = Scan.this_string "(*"
+ − 51
val end_cmt = Scan.this_string "*)"
+ − 52
in
+ − 53
begin_cmt |-- Scan.repeat (Scan.unless end_cmt any) --| end_cmt
+ − 54
>> (enclose "(**" "**)" o implode)
+ − 55
end
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 56
130
+ − 57
val parser = Scan.repeat (scan_cmt || any)
+ − 58
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 59
val scan_all =
168
+ − 60
Scan.finite Symbol.stopper parser >> implode #> fst *}
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 61
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 62
text {*
58
+ − 63
By using @{text "#> fst"} in the last line, the function
+ − 64
@{ML scan_all} retruns a string, instead of the pair a parser would
+ − 65
normally return. For example:
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 66
80
+ − 67
@{ML_response [display,gray]
+ − 68
"let
130
+ − 69
val input1 = (explode \"foo bar\")
+ − 70
val input2 = (explode \"foo (*test*) bar (*test*)\")
+ − 71
in
+ − 72
(scan_all input1, scan_all input2)
+ − 73
end"
+ − 74
"(\"foo bar\", \"foo (**test**) bar (**test**)\")"}
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 75
*}
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 76
314
+ − 77
text {* \solution{ex:dyckhoff} *}
+ − 78
+ − 79
text {*
+ − 80
The axiom rule can be implemented with the function @{ML atac}. The other
+ − 81
rules correspond to the theorems:
+ − 82
+ − 83
\begin{center}
+ − 84
\begin{tabular}{cc}
+ − 85
\begin{tabular}{rl}
+ − 86
$\wedge_R$ & @{thm [source] conjI}\\
+ − 87
$\vee_{R_1}$ & @{thm [source] disjI1}\\
+ − 88
$\vee_{R_2}$ & @{thm [source] disjI2}\\
+ − 89
$\longrightarrow_R$ & @{thm [source] impI}\\
+ − 90
$=_R$ & @{thm [source] iffI}\\
+ − 91
\end{tabular}
+ − 92
&
+ − 93
\begin{tabular}{rl}
+ − 94
$False$ & @{thm [source] FalseE}\\
+ − 95
$\wedge_L$ & @{thm [source] conjE}\\
+ − 96
$\vee_L$ & @{thm [source] disjE}\\
+ − 97
$=_L$ & @{thm [source] iffE}
+ − 98
\end{tabular}
+ − 99
\end{tabular}
+ − 100
\end{center}
+ − 101
+ − 102
For the other rules we need to prove the following lemmas.
+ − 103
*}
+ − 104
+ − 105
lemma impE1:
+ − 106
shows "\<lbrakk>A \<longrightarrow> B; A; B \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+ − 107
by iprover
+ − 108
+ − 109
lemma impE2:
+ − 110
shows "\<lbrakk>(C \<and> D) \<longrightarrow> B; C \<longrightarrow> (D \<longrightarrow>B) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+ − 111
by iprover
+ − 112
+ − 113
lemma impE3:
+ − 114
shows "\<lbrakk>(C \<or> D) \<longrightarrow> B; \<lbrakk>C \<longrightarrow> B; D \<longrightarrow> B\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+ − 115
by iprover
+ − 116
+ − 117
lemma impE4:
+ − 118
shows "\<lbrakk>(C \<longrightarrow> D) \<longrightarrow> B; D \<longrightarrow> B \<Longrightarrow> C \<longrightarrow> D; B \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+ − 119
by iprover
+ − 120
+ − 121
lemma impE5:
+ − 122
shows "\<lbrakk>(C = D) \<longrightarrow> B; (C \<longrightarrow> D) \<longrightarrow> ((D \<longrightarrow> C) \<longrightarrow> B) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+ − 123
by iprover
+ − 124
+ − 125
text {*
+ − 126
Now the tactic which applies a single rule can be implemented
+ − 127
as follows.
+ − 128
*}
+ − 129
+ − 130
ML %linenosgray{*val apply_tac =
+ − 131
let
+ − 132
val intros = [@{thm conjI}, @{thm disjI1}, @{thm disjI2},
+ − 133
@{thm impI}, @{thm iffI}]
+ − 134
val elims = [@{thm FalseE}, @{thm conjE}, @{thm disjE}, @{thm iffE},
+ − 135
@{thm impE2}, @{thm impE3}, @{thm impE4}, @{thm impE5}]
+ − 136
in
+ − 137
atac
+ − 138
ORELSE' resolve_tac intros
+ − 139
ORELSE' eresolve_tac elims
+ − 140
ORELSE' (etac @{thm impE1} THEN' atac)
+ − 141
end*}
+ − 142
+ − 143
text {*
+ − 144
In Line 11 we apply the rule @{thm [source] impE1} in concjunction
+ − 145
with @{ML atac} in order to reduce the number of possibilities that
+ − 146
need to be explored. You can use the tactic as follows.
+ − 147
*}
+ − 148
+ − 149
lemma
+ − 150
shows "((((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> Q) \<longrightarrow> Q"
+ − 151
apply(tactic {* (DEPTH_SOLVE o apply_tac) 1 *})
+ − 152
done
+ − 153
+ − 154
text {*
+ − 155
We can use the tactic to prove or disprove automatically the
+ − 156
de Bruijn formulae from Exercise \ref{ex:debruijn}.
+ − 157
*}
+ − 158
+ − 159
ML{*fun de_bruijn_prove ctxt n =
+ − 160
let
+ − 161
val goal = HOLogic.mk_Trueprop (HOLogic.mk_imp (lhs n n, rhs n))
+ − 162
in
+ − 163
Goal.prove ctxt ["P"] [] goal
+ − 164
(fn _ => (DEPTH_SOLVE o apply_tac) 1)
+ − 165
end*}
+ − 166
+ − 167
text {*
+ − 168
You can use this function to prove de Bruijn formulae.
+ − 169
*}
+ − 170
+ − 171
ML{*de_bruijn_prove @{context} 3 *}
+ − 172
130
+ − 173
text {* \solution{ex:addsimproc} *}
+ − 174
+ − 175
ML{*fun dest_sum term =
+ − 176
case term of
+ − 177
(@{term "(op +):: nat \<Rightarrow> nat \<Rightarrow> nat"} $ t1 $ t2) =>
+ − 178
(snd (HOLogic.dest_number t1), snd (HOLogic.dest_number t2))
+ − 179
| _ => raise TERM ("dest_sum", [term])
+ − 180
+ − 181
fun get_sum_thm ctxt t (n1, n2) =
+ − 182
let
+ − 183
val sum = HOLogic.mk_number @{typ "nat"} (n1 + n2)
132
+ − 184
val goal = Logic.mk_equals (t, sum)
130
+ − 185
in
214
+ − 186
Goal.prove ctxt [] [] goal (K (Arith_Data.arith_tac ctxt 1))
130
+ − 187
end
+ − 188
+ − 189
fun add_sp_aux ss t =
+ − 190
let
+ − 191
val ctxt = Simplifier.the_context ss
+ − 192
val t' = term_of t
+ − 193
in
+ − 194
SOME (get_sum_thm ctxt t' (dest_sum t'))
+ − 195
handle TERM _ => NONE
+ − 196
end*}
+ − 197
+ − 198
text {* The setup for the simproc is *}
+ − 199
177
+ − 200
simproc_setup %gray add_sp ("t1 + t2") = {* K add_sp_aux *}
130
+ − 201
+ − 202
text {* and a test case is the lemma *}
+ − 203
+ − 204
lemma "P (Suc (99 + 1)) ((0 + 0)::nat) (Suc (3 + 3 + 3)) (4 + 1)"
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 205
apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [@{simproc add_sp}]) 1 *})
130
+ − 206
txt {*
+ − 207
where the simproc produces the goal state
+ − 208
+ − 209
\begin{minipage}{\textwidth}
+ − 210
@{subgoals [display]}
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 211
\end{minipage}\bigskip
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 212
*}(*<*)oops(*>*)
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 213
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 214
text {* \solution{ex:addconversion} *}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 215
168
+ − 216
text {*
+ − 217
The following code assumes the function @{ML dest_sum} from the previous
+ − 218
exercise.
+ − 219
*}
+ − 220
166
+ − 221
ML{*fun add_simple_conv ctxt ctrm =
+ − 222
let
+ − 223
val trm = Thm.term_of ctrm
+ − 224
in
+ − 225
get_sum_thm ctxt trm (dest_sum trm)
+ − 226
end
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 227
166
+ − 228
fun add_conv ctxt ctrm =
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 229
(case Thm.term_of ctrm of
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 230
@{term "(op +)::nat \<Rightarrow> nat \<Rightarrow> nat"} $ _ $ _ =>
166
+ − 231
(Conv.binop_conv (add_conv ctxt)
+ − 232
then_conv (Conv.try_conv (add_simple_conv ctxt))) ctrm
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 233
| _ $ _ => Conv.combination_conv
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 234
(add_conv ctxt) (add_conv ctxt) ctrm
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 235
| Abs _ => Conv.abs_conv (fn (_, ctxt) => add_conv ctxt) ctxt ctrm
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 236
| _ => Conv.all_conv ctrm)
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 237
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 238
fun add_tac ctxt = CSUBGOAL (fn (goal, i) =>
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 239
CONVERSION
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 240
(Conv.params_conv ~1 (fn ctxt =>
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 241
(Conv.prems_conv ~1 (add_conv ctxt) then_conv
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 242
Conv.concl_conv ~1 (add_conv ctxt))) ctxt) i)*}
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 243
168
+ − 244
text {*
175
+ − 245
A test case for this conversion is as follows
168
+ − 246
*}
+ − 247
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 248
lemma "P (Suc (99 + 1)) ((0 + 0)::nat) (Suc (3 + 3 + 3)) (4 + 1)"
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 249
apply(tactic {* add_tac @{context} 1 *})?
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 250
txt {*
175
+ − 251
where it produces the goal state
151
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 252
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 253
\begin{minipage}{\textwidth}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 254
@{subgoals [display]}
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 255
\end{minipage}\bigskip
130
+ − 256
*}(*<*)oops(*>*)
+ − 257
172
+ − 258
text {* \solution{ex:addconversion} *}
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 259
172
+ − 260
text {*
174
+ − 261
We use the timing function @{ML timing_wrapper} from Recipe~\ref{rec:timing}.
+ − 262
To measure any difference between the simproc and conversion, we will create
+ − 263
mechanically terms involving additions and then set up a goal to be
+ − 264
simplified. We have to be careful to set up the goal so that
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 265
other parts of the simplifier do not interfere. For this we construct an
174
+ − 266
unprovable goal which, after simplification, we are going to ``prove'' with
192
+ − 267
the help of ``\isacommand{sorry}'', that is the method @{ML SkipProof.cheat_tac}
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 268
174
+ − 269
For constructing test cases, we first define a function that returns a
+ − 270
complete binary tree whose leaves are numbers and the nodes are
+ − 271
additions.
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 272
*}
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 273
172
+ − 274
ML{*fun term_tree n =
+ − 275
let
+ − 276
val count = ref 0;
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 277
172
+ − 278
fun term_tree_aux n =
+ − 279
case n of
+ − 280
0 => (count := !count + 1; HOLogic.mk_number @{typ nat} (!count))
+ − 281
| _ => Const (@{const_name "plus"}, @{typ "nat\<Rightarrow>nat\<Rightarrow>nat"})
+ − 282
$ (term_tree_aux (n - 1)) $ (term_tree_aux (n - 1))
+ − 283
in
+ − 284
term_tree_aux n
+ − 285
end*}
+ − 286
+ − 287
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
+ − 288
This function generates for example:
172
+ − 289
+ − 290
@{ML_response_fake [display,gray]
239
+ − 291
"writeln (Syntax.string_of_term @{context} (term_tree 2))"
172
+ − 292
"(1 + 2) + (3 + 4)"}
+ − 293
174
+ − 294
The next function constructs a goal of the form @{text "P \<dots>"} with a term
+ − 295
produced by @{ML term_tree} filled in.
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 296
*}
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 297
172
+ − 298
ML{*fun goal n = HOLogic.mk_Trueprop (@{term "P::nat\<Rightarrow> bool"} $ (term_tree n))*}
+ − 299
+ − 300
text {*
+ − 301
Note that the goal needs to be wrapped in a @{term "Trueprop"}. Next we define
+ − 302
two tactics, @{text "c_tac"} and @{text "s_tac"}, for the conversion and simproc,
174
+ − 303
respectively. The idea is to first apply the conversion (respectively simproc) and
192
+ − 304
then prove the remaining goal using @{ML "cheat_tac" in SkipProof}.
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 305
*}
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 306
172
+ − 307
ML{*local
225
+ − 308
fun mk_tac tac =
+ − 309
timing_wrapper (EVERY1 [tac, K (SkipProof.cheat_tac @{theory})])
172
+ − 310
in
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 311
val c_tac = mk_tac (add_tac @{context})
174
+ − 312
val s_tac = mk_tac (simp_tac (HOL_basic_ss addsimprocs [@{simproc add_sp}]))
172
+ − 313
end*}
+ − 314
+ − 315
text {*
175
+ − 316
This is all we need to let the conversion run against the simproc:
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 317
*}
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 318
174
+ − 319
ML{*val _ = Goal.prove @{context} [] [] (goal 8) (K c_tac)
172
+ − 320
val _ = Goal.prove @{context} [] [] (goal 8) (K s_tac)*}
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 321
172
+ − 322
text {*
174
+ − 323
If you do the exercise, you can see that both ways of simplifying additions
178
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 324
perform relatively similar with perhaps some advantages for the
174
+ − 325
simproc. That means the simplifier, even if much more complicated than
+ − 326
conversions, is quite efficient for tasks it is designed for. It usually does not
+ − 327
make sense to implement general-purpose rewriting using
172
+ − 328
conversions. Conversions only have clear advantages in special situations:
+ − 329
for example if you need to have control over innermost or outermost
174
+ − 330
rewriting, or when rewriting rules are lead to non-termination.
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 331
*}
130
+ − 332
270
+ − 333
end