author | Norbert Schirmer <norbert.schirmer@web.de> |
Tue, 14 May 2019 11:10:53 +0200 | |
changeset 562 | daf404920ab9 |
parent 544 | 501491d56798 |
child 565 | cecd7a941885 |
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 |
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 |
|
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 |
|
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
9 |
ML %grayML{*fun rev_sum |
447 | 10 |
((p as Const (@{const_name plus}, _)) $ t $ u) = p $ u $ rev_sum t |
271 | 11 |
| rev_sum t = t *} |
12 |
||
293
0a567f923b42
slightly changed exercises about rev_sum
Christian Urban <urbanc@in.tum.de>
parents:
272
diff
changeset
|
13 |
text {* |
316
74f0a06f751f
further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents:
314
diff
changeset
|
14 |
An alternative solution using the function @{ML_ind mk_binop in HOLogic} is: |
293
0a567f923b42
slightly changed exercises about rev_sum
Christian Urban <urbanc@in.tum.de>
parents:
272
diff
changeset
|
15 |
*} |
271 | 16 |
|
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
17 |
ML %grayML{*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
|
18 |
let |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
19 |
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
|
20 |
| dest_sum u = [u] |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
21 |
in |
447 | 22 |
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
|
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 |
|
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
27 |
ML %grayML{*fun make_sum t1 t2 = |
447 | 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 |
|
469
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
30 |
text {* \solution{fun:killqnt} *} |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
31 |
|
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
32 |
ML %linenosgray{*val quantifiers = [@{const_name All}, @{const_name Ex}] |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
33 |
|
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
34 |
fun kill_trivial_quantifiers trm = |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
35 |
let |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
36 |
fun aux t = |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
37 |
case t of |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
38 |
Const (s1, T1) $ Abs (x, T2, t2) => |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
39 |
if member (op =) quantifiers s1 andalso not (loose_bvar1 (t2, 0)) |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
40 |
then incr_boundvars ~1 (aux t2) |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
41 |
else Const (s1, T1) $ Abs (x, T2, aux t2) |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
42 |
| t1 $ t2 => aux t1 $ aux t2 |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
43 |
| Abs (s, T, t') => Abs (s, T, aux t') |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
44 |
| _ => t |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
45 |
in |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
46 |
aux trm |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
47 |
end*} |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
48 |
|
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
49 |
text {* |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
50 |
In line 7 we traverse the term, by first checking whether a term is an |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
51 |
application of a constant with an abstraction. If the constant stands for |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
52 |
a listed quantifier (see Line 1) and the bound variable does not occur |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
53 |
as a loose bound variable in the body, then we delete the quantifier. |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
54 |
For this we have to increase all other dangling de Bruijn indices by |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
55 |
@{text "-1"} to account for the deleted quantifier. An example is |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
56 |
as follows: |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
57 |
|
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
58 |
@{ML_response_fake [display,gray] |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
59 |
"@{prop \"\<forall>x y z. P x = P z\"} |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
60 |
|> kill_trivial_quantifiers |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
61 |
|> pretty_term @{context} |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
62 |
|> pwriteln" |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
63 |
"\<forall>x z. P x = P z"} |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
64 |
*} |
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
65 |
|
7a558c5119b2
added an excercise originally by Jasmin Blanchette
Christian Urban <urbanc@in.tum.de>
parents:
458
diff
changeset
|
66 |
|
446
4c32349b9875
added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents:
441
diff
changeset
|
67 |
|
4c32349b9875
added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents:
441
diff
changeset
|
68 |
text {* \solution{fun:makelist} *} |
4c32349b9875
added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents:
441
diff
changeset
|
69 |
|
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
70 |
ML %grayML{*fun mk_rev_upto i = |
446
4c32349b9875
added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents:
441
diff
changeset
|
71 |
1 upto i |
4c32349b9875
added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents:
441
diff
changeset
|
72 |
|> map (HOLogic.mk_number @{typ int}) |
4c32349b9875
added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents:
441
diff
changeset
|
73 |
|> HOLogic.mk_list @{typ int} |
4c32349b9875
added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents:
441
diff
changeset
|
74 |
|> curry (op $) @{term "rev :: int list \<Rightarrow> int list"}*} |
4c32349b9875
added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents:
441
diff
changeset
|
75 |
|
313
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
76 |
text {* \solution{ex:debruijn} *} |
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
77 |
|
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
78 |
ML %grayML{*fun P n = @{term "P::nat \<Rightarrow> bool"} $ (HOLogic.mk_number @{typ "nat"} n) |
313
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
79 |
|
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
80 |
fun rhs 1 = P 1 |
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
81 |
| rhs n = HOLogic.mk_conj (P n, rhs (n - 1)) |
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
82 |
|
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
83 |
fun lhs 1 n = HOLogic.mk_imp (HOLogic.mk_eq (P 1, P n), rhs n) |
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
84 |
| lhs m n = HOLogic.mk_conj (HOLogic.mk_imp |
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
85 |
(HOLogic.mk_eq (P (m - 1), P m), rhs n), lhs (m - 1) n) |
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
86 |
|
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
87 |
fun de_bruijn n = |
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
88 |
HOLogic.mk_Trueprop (HOLogic.mk_imp (lhs n n, rhs n))*} |
1ca2f41770cc
polished the exercises about constructing terms
Christian Urban <urbanc@in.tum.de>
parents:
312
diff
changeset
|
89 |
|
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents:
47
diff
changeset
|
90 |
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
|
91 |
|
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
92 |
ML %grayML{*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
|
93 |
|
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents:
47
diff
changeset
|
94 |
val scan_cmt = |
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
95 |
let |
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
96 |
val begin_cmt = Scan.this_string "(*" |
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
97 |
val end_cmt = Scan.this_string "*)" |
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
98 |
in |
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
99 |
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
|
100 |
>> (enclose "(**" "**)" o implode) |
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
101 |
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
|
102 |
|
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
103 |
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
|
104 |
|
56
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents:
47
diff
changeset
|
105 |
val scan_all = |
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
106 |
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
|
107 |
|
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents:
47
diff
changeset
|
108 |
text {* |
58 | 109 |
By using @{text "#> fst"} in the last line, the function |
110 |
@{ML scan_all} retruns a string, instead of the pair a parser would |
|
111 |
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
|
112 |
|
80 | 113 |
@{ML_response [display,gray] |
114 |
"let |
|
458
242e81f4d461
updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents:
447
diff
changeset
|
115 |
val input1 = (Symbol.explode \"foo bar\") |
242e81f4d461
updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents:
447
diff
changeset
|
116 |
val input2 = (Symbol.explode \"foo (*test*) bar (*test*)\") |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
117 |
in |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
118 |
(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
|
119 |
end" |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
120 |
"(\"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
|
121 |
*} |
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents:
47
diff
changeset
|
122 |
|
391 | 123 |
text {* \solution{ex:contextfree} *} |
124 |
||
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
125 |
ML %grayML{*datatype expr = |
391 | 126 |
Number of int |
127 |
| Mult of expr * expr |
|
128 |
| Add of expr * expr |
|
314 | 129 |
|
391 | 130 |
fun parse_basic xs = |
426 | 131 |
(Parse.nat >> Number |
132 |
|| Parse.$$$ "(" |-- parse_expr --| Parse.$$$ ")") xs |
|
391 | 133 |
and parse_factor xs = |
426 | 134 |
(parse_basic --| Parse.$$$ "*" -- parse_factor >> Mult |
391 | 135 |
|| parse_basic) xs |
136 |
and parse_expr xs = |
|
426 | 137 |
(parse_factor --| Parse.$$$ "+" -- parse_expr >> Add |
391 | 138 |
|| parse_factor) xs*} |
139 |
||
140 |
||
407 | 141 |
text {* \solution{ex:dyckhoff} *} |
142 |
||
143 |
text {* |
|
562
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
144 |
The axiom rule can be implemented with the function @{ML assume_tac}. The other |
314 | 145 |
rules correspond to the theorems: |
146 |
||
147 |
\begin{center} |
|
148 |
\begin{tabular}{cc} |
|
149 |
\begin{tabular}{rl} |
|
150 |
$\wedge_R$ & @{thm [source] conjI}\\ |
|
151 |
$\vee_{R_1}$ & @{thm [source] disjI1}\\ |
|
152 |
$\vee_{R_2}$ & @{thm [source] disjI2}\\ |
|
153 |
$\longrightarrow_R$ & @{thm [source] impI}\\ |
|
154 |
$=_R$ & @{thm [source] iffI}\\ |
|
155 |
\end{tabular} |
|
156 |
& |
|
157 |
\begin{tabular}{rl} |
|
158 |
$False$ & @{thm [source] FalseE}\\ |
|
159 |
$\wedge_L$ & @{thm [source] conjE}\\ |
|
160 |
$\vee_L$ & @{thm [source] disjE}\\ |
|
161 |
$=_L$ & @{thm [source] iffE} |
|
162 |
\end{tabular} |
|
163 |
\end{tabular} |
|
164 |
\end{center} |
|
165 |
||
166 |
For the other rules we need to prove the following lemmas. |
|
167 |
*} |
|
168 |
||
169 |
lemma impE1: |
|
170 |
shows "\<lbrakk>A \<longrightarrow> B; A; B \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
|
171 |
by iprover |
|
172 |
||
173 |
lemma impE2: |
|
174 |
shows "\<lbrakk>(C \<and> D) \<longrightarrow> B; C \<longrightarrow> (D \<longrightarrow>B) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
|
447 | 175 |
and "\<lbrakk>(C \<or> D) \<longrightarrow> B; \<lbrakk>C \<longrightarrow> B; D \<longrightarrow> B\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
176 |
and "\<lbrakk>(C \<longrightarrow> D) \<longrightarrow> B; D \<longrightarrow> B \<Longrightarrow> C \<longrightarrow> D; B \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
|
177 |
and "\<lbrakk>(C = D) \<longrightarrow> B; (C \<longrightarrow> D) \<longrightarrow> ((D \<longrightarrow> C) \<longrightarrow> B) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
|
178 |
by iprover+ |
|
314 | 179 |
|
180 |
text {* |
|
181 |
Now the tactic which applies a single rule can be implemented |
|
182 |
as follows. |
|
183 |
*} |
|
184 |
||
562
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
185 |
ML %linenosgray{*fun apply_tac ctxt = |
314 | 186 |
let |
447 | 187 |
val intros = @{thms conjI disjI1 disjI2 impI iffI} |
188 |
val elims = @{thms FalseE conjE disjE iffE impE2} |
|
314 | 189 |
in |
562
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
190 |
assume_tac ctxt |
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
191 |
ORELSE' resolve_tac ctxt intros |
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
192 |
ORELSE' eresolve_tac ctxt elims |
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
193 |
ORELSE' (eresolve_tac ctxt [@{thm impE1}] THEN' assume_tac ctxt) |
314 | 194 |
end*} |
195 |
||
196 |
text {* |
|
197 |
In Line 11 we apply the rule @{thm [source] impE1} in concjunction |
|
562
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
198 |
with @{ML assume_tac} in order to reduce the number of possibilities that |
314 | 199 |
need to be explored. You can use the tactic as follows. |
200 |
*} |
|
201 |
||
202 |
lemma |
|
203 |
shows "((((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> Q) \<longrightarrow> Q" |
|
562
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
204 |
apply(tactic {* (DEPTH_SOLVE o apply_tac @{context}) 1 *}) |
314 | 205 |
done |
206 |
||
207 |
text {* |
|
208 |
We can use the tactic to prove or disprove automatically the |
|
209 |
de Bruijn formulae from Exercise \ref{ex:debruijn}. |
|
210 |
*} |
|
211 |
||
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
212 |
ML %grayML{*fun de_bruijn_prove ctxt n = |
314 | 213 |
let |
214 |
val goal = HOLogic.mk_Trueprop (HOLogic.mk_imp (lhs n n, rhs n)) |
|
215 |
in |
|
216 |
Goal.prove ctxt ["P"] [] goal |
|
562
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
217 |
(fn _ => (DEPTH_SOLVE o apply_tac ctxt) 1) |
314 | 218 |
end*} |
219 |
||
220 |
text {* |
|
221 |
You can use this function to prove de Bruijn formulae. |
|
222 |
*} |
|
223 |
||
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
224 |
ML %grayML{*de_bruijn_prove @{context} 3 *} |
314 | 225 |
|
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
226 |
text {* \solution{ex:addsimproc} *} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
227 |
|
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
228 |
ML %grayML{*fun dest_sum term = |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
229 |
case term of |
562
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
230 |
(@{term "(+):: nat \<Rightarrow> nat \<Rightarrow> nat"} $ t1 $ t2) => |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
231 |
(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
|
232 |
| _ => 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
|
233 |
|
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
234 |
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
|
235 |
let |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
236 |
val sum = HOLogic.mk_number @{typ "nat"} (n1 + n2) |
132 | 237 |
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
|
238 |
in |
214
7e04dc2368b0
updated to latest Isabelle
Christian Urban <urbanc@in.tum.de>
parents:
192
diff
changeset
|
239 |
Goal.prove ctxt [] [] goal (K (Arith_Data.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
|
240 |
end |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
241 |
|
544
501491d56798
updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
517
diff
changeset
|
242 |
fun add_sp_aux ctxt t = |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
243 |
let |
562
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
244 |
val t' = Thm.term_of t |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
245 |
in |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
246 |
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
|
247 |
handle TERM _ => NONE |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
248 |
end*} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
249 |
|
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
250 |
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
|
251 |
|
177 | 252 |
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
|
253 |
|
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
254 |
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
|
255 |
|
405
f8d020bbc2c0
improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents:
402
diff
changeset
|
256 |
lemma "P (Suc (99 + 1)) ((0 + 0)::nat) (Suc (3 + 3 + 3)) ((4 + 1)::nat)" |
544
501491d56798
updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
517
diff
changeset
|
257 |
apply(tactic {* simp_tac (put_simpset HOL_basic_ss @{context} 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
|
258 |
txt {* |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
259 |
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
|
260 |
|
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
261 |
\begin{minipage}{\textwidth} |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
262 |
@{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
|
263 |
\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
|
264 |
*}(*<*)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
|
265 |
|
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
|
266 |
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
|
267 |
|
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
268 |
text {* |
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
269 |
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
|
270 |
exercise. |
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
271 |
*} |
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
272 |
|
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
273 |
ML %grayML{*fun add_simple_conv ctxt ctrm = |
166
00d153e32a53
improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents:
158
diff
changeset
|
274 |
let |
405
f8d020bbc2c0
improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents:
402
diff
changeset
|
275 |
val trm = Thm.term_of ctrm |
f8d020bbc2c0
improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents:
402
diff
changeset
|
276 |
in |
f8d020bbc2c0
improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents:
402
diff
changeset
|
277 |
case trm of |
562
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
278 |
@{term "(+)::nat \<Rightarrow> nat \<Rightarrow> nat"} $ _ $ _ => |
405
f8d020bbc2c0
improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents:
402
diff
changeset
|
279 |
get_sum_thm ctxt trm (dest_sum trm) |
f8d020bbc2c0
improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents:
402
diff
changeset
|
280 |
| _ => Conv.all_conv ctrm |
166
00d153e32a53
improvments to the solutions suggested by Sacha B?hme
Christian Urban <urbanc@in.tum.de>
parents:
158
diff
changeset
|
281 |
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
|
282 |
|
424 | 283 |
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>
parents:
132
diff
changeset
|
284 |
|
410 | 285 |
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>
parents:
132
diff
changeset
|
286 |
|
168
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
287 |
text {* |
175 | 288 |
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
|
289 |
*} |
009ca4807baa
polished somewhat the recipes and solutions
Christian Urban <urbanc@in.tum.de>
parents:
167
diff
changeset
|
290 |
|
405
f8d020bbc2c0
improved section on conversions
Christian Urban <urbanc@in.tum.de>
parents:
402
diff
changeset
|
291 |
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>
parents:
178
diff
changeset
|
292 |
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
|
293 |
txt {* |
175 | 294 |
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
|
295 |
|
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
|
296 |
\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
|
297 |
@{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
|
298 |
\end{minipage}\bigskip |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
299 |
*}(*<*)oops(*>*) |
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
300 |
|
407 | 301 |
text {* \solution{ex:compare} *} |
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents:
166
diff
changeset
|
302 |
|
172
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
303 |
text {* |
174 | 304 |
We use the timing function @{ML timing_wrapper} from Recipe~\ref{rec:timing}. |
305 |
To measure any difference between the simproc and conversion, we will create |
|
306 |
mechanically terms involving additions and then set up a goal to be |
|
307 |
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
|
308 |
other parts of the simplifier do not interfere. For this we construct an |
174 | 309 |
unprovable goal which, after simplification, we are going to ``prove'' with |
351 | 310 |
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>
parents:
166
diff
changeset
|
311 |
|
174 | 312 |
For constructing test cases, we first define a function that returns a |
313 |
complete binary tree whose leaves are numbers and the nodes are |
|
314 |
additions. |
|
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents:
166
diff
changeset
|
315 |
*} |
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents:
166
diff
changeset
|
316 |
|
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
317 |
ML %grayML{*fun term_tree n = |
172
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
318 |
let |
328
c0cae24b9d46
updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents:
316
diff
changeset
|
319 |
val count = Unsynchronized.ref 0; |
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents:
166
diff
changeset
|
320 |
|
172
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
321 |
fun term_tree_aux n = |
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
322 |
case n of |
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
323 |
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
|
324 |
| _ => 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
|
325 |
$ (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
|
326 |
in |
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
327 |
term_tree_aux n |
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
328 |
end*} |
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
329 |
|
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
330 |
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
|
331 |
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
|
332 |
|
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
333 |
@{ML_response_fake [display,gray] |
440
a0b280dd4bc7
partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents:
426
diff
changeset
|
334 |
"pwriteln (pretty_term @{context} (term_tree 2))" |
172
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
335 |
"(1 + 2) + (3 + 4)"} |
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
336 |
|
174 | 337 |
The next function constructs a goal of the form @{text "P \<dots>"} with a term |
338 |
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
|
339 |
*} |
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents:
166
diff
changeset
|
340 |
|
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
341 |
ML %grayML{*fun goal n = HOLogic.mk_Trueprop (@{term "P::nat\<Rightarrow> bool"} $ (term_tree n))*} |
172
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
342 |
|
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
343 |
text {* |
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
344 |
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
|
345 |
two tactics, @{text "c_tac"} and @{text "s_tac"}, for the conversion and simproc, |
174 | 346 |
respectively. The idea is to first apply the conversion (respectively simproc) and |
351 | 347 |
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>
parents:
166
diff
changeset
|
348 |
*} |
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents:
166
diff
changeset
|
349 |
|
544
501491d56798
updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
517
diff
changeset
|
350 |
ML Skip_Proof.cheat_tac |
501491d56798
updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
517
diff
changeset
|
351 |
|
517
d8c376662bb4
removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents:
469
diff
changeset
|
352 |
ML %grayML{*local |
562
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
353 |
fun mk_tac ctxt tac = |
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
354 |
timing_wrapper (EVERY1 [tac, Skip_Proof.cheat_tac ctxt]) |
172
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
355 |
in |
562
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
356 |
fun c_tac ctxt = mk_tac ctxt (add_tac ctxt) |
daf404920ab9
Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents:
544
diff
changeset
|
357 |
fun s_tac ctxt = mk_tac ctxt (simp_tac |
544
501491d56798
updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
517
diff
changeset
|
358 |
(put_simpset HOL_basic_ss ctxt addsimprocs [@{simproc add_sp}])) |
172
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
359 |
end*} |
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
360 |
|
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
361 |
text {* |
175 | 362 |
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
|
363 |
*} |
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents:
166
diff
changeset
|
364 |
|
544
501491d56798
updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
517
diff
changeset
|
365 |
ML %grayML{*val _ = Goal.prove @{context} [] [] (goal 8) (fn {context, ...} => c_tac context) |
501491d56798
updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
517
diff
changeset
|
366 |
val _ = Goal.prove @{context} [] [] (goal 8) (fn {context, ...} => s_tac context)*} |
167
3e30ea95c7aa
added temporarily some timing test about conversions and simprocs
Christian Urban <urbanc@in.tum.de>
parents:
166
diff
changeset
|
367 |
|
172
ec47352e99c2
improved the solution for the simproc/conversion exercise
Christian Urban <urbanc@in.tum.de>
parents:
168
diff
changeset
|
368 |
text {* |
174 | 369 |
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
|
370 |
perform relatively similar with perhaps some advantages for the |
174 | 371 |
simproc. That means the simplifier, even if much more complicated than |
372 |
conversions, is quite efficient for tasks it is designed for. It usually does not |
|
373 |
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
|
374 |
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
|
375 |
for example if you need to have control over innermost or outermost |
174 | 376 |
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
|
377 |
*} |
130
a21d7b300616
polished the section about simprocs and added an exercise
Christian Urban <urbanc@in.tum.de>
parents:
80
diff
changeset
|
378 |
|
270
bfcabed9079e
included an alternative solution for the "rev_sum" example as a comment
griff
parents:
225
diff
changeset
|
379 |
end |