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