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