author | Christian Urban <urbanc@in.tum.de> |
Mon, 23 Feb 2009 17:08:15 +0000 | |
changeset 132 | 2d9198bcb850 |
parent 130 | a21d7b300616 |
child 151 | 7e0bf13bf743 |
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 |
25
e2f9f94b26d4
Antiquotation setup is now contained in theory Base.
berghofe
parents:
15
diff
changeset
|
2 |
imports Base |
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 |
|
9da9ba2b095b
added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
chapter {* Solutions to Most Exercises *} |
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 = |
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents:
47
diff
changeset
|
27 |
let |
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents:
47
diff
changeset
|
28 |
val begin_cmt = Scan.this_string "(*" |
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents:
47
diff
changeset
|
29 |
val end_cmt = Scan.this_string "*)" |
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents:
47
diff
changeset
|
30 |
in |
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents:
47
diff
changeset
|
31 |
begin_cmt |-- Scan.repeat (Scan.unless end_cmt any) --| end_cmt |
58 | 32 |
>> (enclose "(**" "**)" o implode) |
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents:
47
diff
changeset
|
33 |
end |
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 = |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
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 |
|
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
82 |
simproc_setup add_sp ("t1 + t2") = {* K add_sp_aux *} |
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)" |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
87 |
apply(tactic {* simp_tac (HOL_ss addsimprocs [@{simproc add_sp}]) 1 *}) |
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]} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
93 |
\end{minipage} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
94 |
*}(*<*)oops(*>*) |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
95 |
|
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
96 |
|
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
97 |
|
15
9da9ba2b095b
added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
98 |
end |