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 {*
+ − 14
An alternative solution using the function @{ML [index] mk_binop in HOLogic} is:
+ − 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
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 30
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
+ − 31
130
+ − 32
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
+ − 33
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 34
val scan_cmt =
168
+ − 35
let
+ − 36
val begin_cmt = Scan.this_string "(*"
+ − 37
val end_cmt = Scan.this_string "*)"
+ − 38
in
+ − 39
begin_cmt |-- Scan.repeat (Scan.unless end_cmt any) --| end_cmt
+ − 40
>> (enclose "(**" "**)" o implode)
+ − 41
end
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 42
130
+ − 43
val parser = Scan.repeat (scan_cmt || any)
+ − 44
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 45
val scan_all =
168
+ − 46
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
+ − 47
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 48
text {*
58
+ − 49
By using @{text "#> fst"} in the last line, the function
+ − 50
@{ML scan_all} retruns a string, instead of the pair a parser would
+ − 51
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
+ − 52
80
+ − 53
@{ML_response [display,gray]
+ − 54
"let
130
+ − 55
val input1 = (explode \"foo bar\")
+ − 56
val input2 = (explode \"foo (*test*) bar (*test*)\")
+ − 57
in
+ − 58
(scan_all input1, scan_all input2)
+ − 59
end"
+ − 60
"(\"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
+ − 61
*}
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 62
130
+ − 63
text {* \solution{ex:addsimproc} *}
+ − 64
+ − 65
ML{*fun dest_sum term =
+ − 66
case term of
+ − 67
(@{term "(op +):: nat \<Rightarrow> nat \<Rightarrow> nat"} $ t1 $ t2) =>
+ − 68
(snd (HOLogic.dest_number t1), snd (HOLogic.dest_number t2))
+ − 69
| _ => raise TERM ("dest_sum", [term])
+ − 70
+ − 71
fun get_sum_thm ctxt t (n1, n2) =
+ − 72
let
+ − 73
val sum = HOLogic.mk_number @{typ "nat"} (n1 + n2)
132
+ − 74
val goal = Logic.mk_equals (t, sum)
130
+ − 75
in
214
+ − 76
Goal.prove ctxt [] [] goal (K (Arith_Data.arith_tac ctxt 1))
130
+ − 77
end
+ − 78
+ − 79
fun add_sp_aux ss t =
+ − 80
let
+ − 81
val ctxt = Simplifier.the_context ss
+ − 82
val t' = term_of t
+ − 83
in
+ − 84
SOME (get_sum_thm ctxt t' (dest_sum t'))
+ − 85
handle TERM _ => NONE
+ − 86
end*}
+ − 87
+ − 88
text {* The setup for the simproc is *}
+ − 89
177
+ − 90
simproc_setup %gray add_sp ("t1 + t2") = {* K add_sp_aux *}
130
+ − 91
+ − 92
text {* and a test case is the lemma *}
+ − 93
+ − 94
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
+ − 95
apply(tactic {* simp_tac (HOL_basic_ss addsimprocs [@{simproc add_sp}]) 1 *})
130
+ − 96
txt {*
+ − 97
where the simproc produces the goal state
+ − 98
+ − 99
\begin{minipage}{\textwidth}
+ − 100
@{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
+ − 101
\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
+ − 102
*}(*<*)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
+ − 103
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
+ − 104
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
+ − 105
168
+ − 106
text {*
+ − 107
The following code assumes the function @{ML dest_sum} from the previous
+ − 108
exercise.
+ − 109
*}
+ − 110
166
+ − 111
ML{*fun add_simple_conv ctxt ctrm =
+ − 112
let
+ − 113
val trm = Thm.term_of ctrm
+ − 114
in
+ − 115
get_sum_thm ctxt trm (dest_sum trm)
+ − 116
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
+ − 117
166
+ − 118
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
+ − 119
(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
+ − 120
@{term "(op +)::nat \<Rightarrow> nat \<Rightarrow> nat"} $ _ $ _ =>
166
+ − 121
(Conv.binop_conv (add_conv ctxt)
+ − 122
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
+ − 123
| _ $ _ => 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
+ − 124
(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
+ − 125
| 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
+ − 126
| _ => 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
+ − 127
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 128
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
+ − 129
CONVERSION
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 130
(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
+ − 131
(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
+ − 132
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
+ − 133
168
+ − 134
text {*
175
+ − 135
A test case for this conversion is as follows
168
+ − 136
*}
+ − 137
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
+ − 138
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
+ − 139
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
+ − 140
txt {*
175
+ − 141
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
+ − 142
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
+ − 143
\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
+ − 144
@{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
+ − 145
\end{minipage}\bigskip
130
+ − 146
*}(*<*)oops(*>*)
+ − 147
172
+ − 148
text {* \solution{ex:addconversion} *}
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 149
172
+ − 150
text {*
174
+ − 151
We use the timing function @{ML timing_wrapper} from Recipe~\ref{rec:timing}.
+ − 152
To measure any difference between the simproc and conversion, we will create
+ − 153
mechanically terms involving additions and then set up a goal to be
+ − 154
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
+ − 155
other parts of the simplifier do not interfere. For this we construct an
174
+ − 156
unprovable goal which, after simplification, we are going to ``prove'' with
192
+ − 157
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
+ − 158
174
+ − 159
For constructing test cases, we first define a function that returns a
+ − 160
complete binary tree whose leaves are numbers and the nodes are
+ − 161
additions.
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 162
*}
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 163
172
+ − 164
ML{*fun term_tree n =
+ − 165
let
+ − 166
val count = ref 0;
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 167
172
+ − 168
fun term_tree_aux n =
+ − 169
case n of
+ − 170
0 => (count := !count + 1; HOLogic.mk_number @{typ nat} (!count))
+ − 171
| _ => Const (@{const_name "plus"}, @{typ "nat\<Rightarrow>nat\<Rightarrow>nat"})
+ − 172
$ (term_tree_aux (n - 1)) $ (term_tree_aux (n - 1))
+ − 173
in
+ − 174
term_tree_aux n
+ − 175
end*}
+ − 176
+ − 177
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
+ − 178
This function generates for example:
172
+ − 179
+ − 180
@{ML_response_fake [display,gray]
239
+ − 181
"writeln (Syntax.string_of_term @{context} (term_tree 2))"
172
+ − 182
"(1 + 2) + (3 + 4)"}
+ − 183
174
+ − 184
The next function constructs a goal of the form @{text "P \<dots>"} with a term
+ − 185
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
+ − 186
*}
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 187
172
+ − 188
ML{*fun goal n = HOLogic.mk_Trueprop (@{term "P::nat\<Rightarrow> bool"} $ (term_tree n))*}
+ − 189
+ − 190
text {*
+ − 191
Note that the goal needs to be wrapped in a @{term "Trueprop"}. Next we define
+ − 192
two tactics, @{text "c_tac"} and @{text "s_tac"}, for the conversion and simproc,
174
+ − 193
respectively. The idea is to first apply the conversion (respectively simproc) and
192
+ − 194
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
+ − 195
*}
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 196
172
+ − 197
ML{*local
225
+ − 198
fun mk_tac tac =
+ − 199
timing_wrapper (EVERY1 [tac, K (SkipProof.cheat_tac @{theory})])
172
+ − 200
in
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 201
val c_tac = mk_tac (add_tac @{context})
174
+ − 202
val s_tac = mk_tac (simp_tac (HOL_basic_ss addsimprocs [@{simproc add_sp}]))
172
+ − 203
end*}
+ − 204
+ − 205
text {*
175
+ − 206
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
+ − 207
*}
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 208
174
+ − 209
ML{*val _ = Goal.prove @{context} [] [] (goal 8) (K c_tac)
172
+ − 210
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
+ − 211
172
+ − 212
text {*
174
+ − 213
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
+ − 214
perform relatively similar with perhaps some advantages for the
174
+ − 215
simproc. That means the simplifier, even if much more complicated than
+ − 216
conversions, is quite efficient for tasks it is designed for. It usually does not
+ − 217
make sense to implement general-purpose rewriting using
172
+ − 218
conversions. Conversions only have clear advantages in special situations:
+ − 219
for example if you need to have control over innermost or outermost
174
+ − 220
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
+ − 221
*}
130
+ − 222
270
+ − 223
end