author | haftmann |
Mon, 12 Jul 2010 08:58:27 +0200 | |
changeset 37766 | a779f463bae4 |
parent 36844 | 5f9385ecc1a7 |
child 38991 | 0e2798f30087 |
permissions | -rw-r--r-- |
23449 | 1 |
(* Title: HOL/MetisTest/BT.thy |
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
36487
50fd056cc3ce
insert a nice proof found by Vampire, which demonstrates the use of "let" in Isar proofs
blanchet
parents:
36484
diff
changeset
|
3 |
Author: Jasmin Blanchette, TU Muenchen |
23449 | 4 |
|
5 |
Testing the metis method |
|
6 |
*) |
|
7 |
||
8 |
header {* Binary trees *} |
|
9 |
||
27104
791607529f6d
rep_datatype command now takes list of constructors as input arguments
haftmann
parents:
26312
diff
changeset
|
10 |
theory BT |
791607529f6d
rep_datatype command now takes list of constructors as input arguments
haftmann
parents:
26312
diff
changeset
|
11 |
imports Main |
791607529f6d
rep_datatype command now takes list of constructors as input arguments
haftmann
parents:
26312
diff
changeset
|
12 |
begin |
23449 | 13 |
|
14 |
datatype 'a bt = |
|
15 |
Lf |
|
16 |
| Br 'a "'a bt" "'a bt" |
|
17 |
||
18 |
consts |
|
19 |
n_nodes :: "'a bt => nat" |
|
20 |
n_leaves :: "'a bt => nat" |
|
21 |
depth :: "'a bt => nat" |
|
22 |
reflect :: "'a bt => 'a bt" |
|
23 |
bt_map :: "('a => 'b) => ('a bt => 'b bt)" |
|
24 |
preorder :: "'a bt => 'a list" |
|
25 |
inorder :: "'a bt => 'a list" |
|
26 |
postorder :: "'a bt => 'a list" |
|
27 |
appnd :: "'a bt => 'a bt => 'a bt" |
|
28 |
||
29 |
primrec |
|
30 |
"n_nodes Lf = 0" |
|
31 |
"n_nodes (Br a t1 t2) = Suc (n_nodes t1 + n_nodes t2)" |
|
32 |
||
33 |
primrec |
|
34 |
"n_leaves Lf = Suc 0" |
|
35 |
"n_leaves (Br a t1 t2) = n_leaves t1 + n_leaves t2" |
|
36 |
||
37 |
primrec |
|
38 |
"depth Lf = 0" |
|
39 |
"depth (Br a t1 t2) = Suc (max (depth t1) (depth t2))" |
|
40 |
||
41 |
primrec |
|
42 |
"reflect Lf = Lf" |
|
43 |
"reflect (Br a t1 t2) = Br a (reflect t2) (reflect t1)" |
|
44 |
||
45 |
primrec |
|
46 |
"bt_map f Lf = Lf" |
|
47 |
"bt_map f (Br a t1 t2) = Br (f a) (bt_map f t1) (bt_map f t2)" |
|
48 |
||
49 |
primrec |
|
50 |
"preorder Lf = []" |
|
51 |
"preorder (Br a t1 t2) = [a] @ (preorder t1) @ (preorder t2)" |
|
52 |
||
53 |
primrec |
|
54 |
"inorder Lf = []" |
|
55 |
"inorder (Br a t1 t2) = (inorder t1) @ [a] @ (inorder t2)" |
|
56 |
||
57 |
primrec |
|
58 |
"postorder Lf = []" |
|
59 |
"postorder (Br a t1 t2) = (postorder t1) @ (postorder t2) @ [a]" |
|
60 |
||
61 |
primrec |
|
62 |
"appnd Lf t = t" |
|
63 |
"appnd (Br a t1 t2) t = Br a (appnd t1 t) (appnd t2 t)" |
|
64 |
||
65 |
||
66 |
text {* \medskip BT simplification *} |
|
67 |
||
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
68 |
declare [[ atp_problem_prefix = "BT__n_leaves_reflect" ]] |
36484 | 69 |
|
23449 | 70 |
lemma n_leaves_reflect: "n_leaves (reflect t) = n_leaves t" |
36484 | 71 |
proof (induct t) |
36487
50fd056cc3ce
insert a nice proof found by Vampire, which demonstrates the use of "let" in Isar proofs
blanchet
parents:
36484
diff
changeset
|
72 |
case Lf thus ?case |
50fd056cc3ce
insert a nice proof found by Vampire, which demonstrates the use of "let" in Isar proofs
blanchet
parents:
36484
diff
changeset
|
73 |
proof - |
50fd056cc3ce
insert a nice proof found by Vampire, which demonstrates the use of "let" in Isar proofs
blanchet
parents:
36484
diff
changeset
|
74 |
let "?p\<^isub>1 x\<^isub>1" = "x\<^isub>1 \<noteq> n_leaves (reflect (Lf::'a bt))" |
50fd056cc3ce
insert a nice proof found by Vampire, which demonstrates the use of "let" in Isar proofs
blanchet
parents:
36484
diff
changeset
|
75 |
have "\<not> ?p\<^isub>1 (Suc 0)" by (metis reflect.simps(1) n_leaves.simps(1)) |
50fd056cc3ce
insert a nice proof found by Vampire, which demonstrates the use of "let" in Isar proofs
blanchet
parents:
36484
diff
changeset
|
76 |
hence "\<not> ?p\<^isub>1 (n_leaves (Lf::'a bt))" by (metis n_leaves.simps(1)) |
50fd056cc3ce
insert a nice proof found by Vampire, which demonstrates the use of "let" in Isar proofs
blanchet
parents:
36484
diff
changeset
|
77 |
thus "n_leaves (reflect (Lf::'a bt)) = n_leaves (Lf::'a bt)" by metis |
50fd056cc3ce
insert a nice proof found by Vampire, which demonstrates the use of "let" in Isar proofs
blanchet
parents:
36484
diff
changeset
|
78 |
qed |
36484 | 79 |
next |
80 |
case (Br a t1 t2) thus ?case |
|
36487
50fd056cc3ce
insert a nice proof found by Vampire, which demonstrates the use of "let" in Isar proofs
blanchet
parents:
36484
diff
changeset
|
81 |
by (metis n_leaves.simps(2) nat_add_commute reflect.simps(2)) |
36484 | 82 |
qed |
23449 | 83 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
84 |
declare [[ atp_problem_prefix = "BT__n_nodes_reflect" ]] |
36484 | 85 |
|
23449 | 86 |
lemma n_nodes_reflect: "n_nodes (reflect t) = n_nodes t" |
36484 | 87 |
proof (induct t) |
88 |
case Lf thus ?case by (metis reflect.simps(1)) |
|
89 |
next |
|
90 |
case (Br a t1 t2) thus ?case |
|
36844 | 91 |
by (metis add_commute n_nodes.simps(2) reflect.simps(2)) |
36484 | 92 |
qed |
23449 | 93 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
94 |
declare [[ atp_problem_prefix = "BT__depth_reflect" ]] |
36484 | 95 |
|
23449 | 96 |
lemma depth_reflect: "depth (reflect t) = depth t" |
36484 | 97 |
apply (induct t) |
98 |
apply (metis depth.simps(1) reflect.simps(1)) |
|
99 |
by (metis depth.simps(2) min_max.inf_sup_aci(5) reflect.simps(2)) |
|
23449 | 100 |
|
101 |
text {* |
|
36484 | 102 |
The famous relationship between the numbers of leaves and nodes. |
23449 | 103 |
*} |
104 |
||
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
105 |
declare [[ atp_problem_prefix = "BT__n_leaves_nodes" ]] |
36484 | 106 |
|
23449 | 107 |
lemma n_leaves_nodes: "n_leaves t = Suc (n_nodes t)" |
36484 | 108 |
apply (induct t) |
109 |
apply (metis n_leaves.simps(1) n_nodes.simps(1)) |
|
110 |
by auto |
|
23449 | 111 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
112 |
declare [[ atp_problem_prefix = "BT__reflect_reflect_ident" ]] |
36484 | 113 |
|
23449 | 114 |
lemma reflect_reflect_ident: "reflect (reflect t) = t" |
36484 | 115 |
apply (induct t) |
116 |
apply (metis reflect.simps(1)) |
|
117 |
proof - |
|
118 |
fix a :: 'a and t1 :: "'a bt" and t2 :: "'a bt" |
|
119 |
assume A1: "reflect (reflect t1) = t1" |
|
120 |
assume A2: "reflect (reflect t2) = t2" |
|
121 |
have "\<And>V U. reflect (Br U V (reflect t1)) = Br U t1 (reflect V)" |
|
122 |
using A1 by (metis reflect.simps(2)) |
|
123 |
hence "\<And>V U. Br U t1 (reflect (reflect V)) = reflect (reflect (Br U t1 V))" |
|
124 |
by (metis reflect.simps(2)) |
|
125 |
hence "\<And>U. reflect (reflect (Br U t1 t2)) = Br U t1 t2" |
|
126 |
using A2 by metis |
|
127 |
thus "reflect (reflect (Br a t1 t2)) = Br a t1 t2" by blast |
|
128 |
qed |
|
23449 | 129 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
130 |
declare [[ atp_problem_prefix = "BT__bt_map_ident" ]] |
36484 | 131 |
|
23449 | 132 |
lemma bt_map_ident: "bt_map (%x. x) = (%y. y)" |
133 |
apply (rule ext) |
|
134 |
apply (induct_tac y) |
|
36484 | 135 |
apply (metis bt_map.simps(1)) |
36571 | 136 |
by (metis bt_map.simps(2)) |
23449 | 137 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
138 |
declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]] |
36484 | 139 |
|
23449 | 140 |
lemma bt_map_appnd: "bt_map f (appnd t u) = appnd (bt_map f t) (bt_map f u)" |
141 |
apply (induct t) |
|
36484 | 142 |
apply (metis appnd.simps(1) bt_map.simps(1)) |
143 |
by (metis appnd.simps(2) bt_map.simps(2)) |
|
23449 | 144 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
145 |
declare [[ atp_problem_prefix = "BT__bt_map_compose" ]] |
36484 | 146 |
|
23449 | 147 |
lemma bt_map_compose: "bt_map (f o g) t = bt_map f (bt_map g t)" |
36484 | 148 |
apply (induct t) |
149 |
apply (metis bt_map.simps(1)) |
|
150 |
by (metis bt_map.simps(2) o_eq_dest_lhs) |
|
23449 | 151 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
152 |
declare [[ atp_problem_prefix = "BT__bt_map_reflect" ]] |
36484 | 153 |
|
23449 | 154 |
lemma bt_map_reflect: "bt_map f (reflect t) = reflect (bt_map f t)" |
36484 | 155 |
apply (induct t) |
156 |
apply (metis bt_map.simps(1) reflect.simps(1)) |
|
157 |
by (metis bt_map.simps(2) reflect.simps(2)) |
|
23449 | 158 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
159 |
declare [[ atp_problem_prefix = "BT__preorder_bt_map" ]] |
36484 | 160 |
|
23449 | 161 |
lemma preorder_bt_map: "preorder (bt_map f t) = map f (preorder t)" |
36484 | 162 |
apply (induct t) |
163 |
apply (metis bt_map.simps(1) map.simps(1) preorder.simps(1)) |
|
164 |
by simp |
|
23449 | 165 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
166 |
declare [[ atp_problem_prefix = "BT__inorder_bt_map" ]] |
36484 | 167 |
|
23449 | 168 |
lemma inorder_bt_map: "inorder (bt_map f t) = map f (inorder t)" |
36484 | 169 |
proof (induct t) |
170 |
case Lf thus ?case |
|
171 |
proof - |
|
172 |
have "map f [] = []" by (metis map.simps(1)) |
|
173 |
hence "map f [] = inorder Lf" by (metis inorder.simps(1)) |
|
174 |
hence "inorder (bt_map f Lf) = map f []" by (metis bt_map.simps(1)) |
|
175 |
thus "inorder (bt_map f Lf) = map f (inorder Lf)" by (metis inorder.simps(1)) |
|
176 |
qed |
|
177 |
next |
|
178 |
case (Br a t1 t2) thus ?case by simp |
|
179 |
qed |
|
23449 | 180 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
181 |
declare [[ atp_problem_prefix = "BT__postorder_bt_map" ]] |
36484 | 182 |
|
23449 | 183 |
lemma postorder_bt_map: "postorder (bt_map f t) = map f (postorder t)" |
36484 | 184 |
apply (induct t) |
185 |
apply (metis Nil_is_map_conv bt_map.simps(1) postorder.simps(1)) |
|
186 |
by simp |
|
23449 | 187 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
188 |
declare [[ atp_problem_prefix = "BT__depth_bt_map" ]] |
36484 | 189 |
|
23449 | 190 |
lemma depth_bt_map [simp]: "depth (bt_map f t) = depth t" |
36484 | 191 |
apply (induct t) |
192 |
apply (metis bt_map.simps(1) depth.simps(1)) |
|
193 |
by simp |
|
23449 | 194 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
195 |
declare [[ atp_problem_prefix = "BT__n_leaves_bt_map" ]] |
36484 | 196 |
|
23449 | 197 |
lemma n_leaves_bt_map [simp]: "n_leaves (bt_map f t) = n_leaves t" |
36484 | 198 |
apply (induct t) |
199 |
apply (metis bt_map.simps(1) n_leaves.simps(1)) |
|
200 |
proof - |
|
201 |
fix a :: 'b and t1 :: "'b bt" and t2 :: "'b bt" |
|
202 |
assume A1: "n_leaves (bt_map f t1) = n_leaves t1" |
|
203 |
assume A2: "n_leaves (bt_map f t2) = n_leaves t2" |
|
204 |
have "\<And>V U. n_leaves (Br U (bt_map f t1) V) = n_leaves t1 + n_leaves V" |
|
205 |
using A1 by (metis n_leaves.simps(2)) |
|
206 |
hence "\<And>V U. n_leaves (bt_map f (Br U t1 V)) = n_leaves t1 + n_leaves (bt_map f V)" |
|
207 |
by (metis bt_map.simps(2)) |
|
208 |
hence F1: "\<And>U. n_leaves (bt_map f (Br U t1 t2)) = n_leaves t1 + n_leaves t2" |
|
209 |
using A2 by metis |
|
210 |
have "n_leaves t1 + n_leaves t2 = n_leaves (Br a t1 t2)" |
|
211 |
by (metis n_leaves.simps(2)) |
|
212 |
thus "n_leaves (bt_map f (Br a t1 t2)) = n_leaves (Br a t1 t2)" |
|
213 |
using F1 by metis |
|
214 |
qed |
|
23449 | 215 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
216 |
declare [[ atp_problem_prefix = "BT__preorder_reflect" ]] |
36484 | 217 |
|
23449 | 218 |
lemma preorder_reflect: "preorder (reflect t) = rev (postorder t)" |
36484 | 219 |
apply (induct t) |
220 |
apply (metis Nil_is_rev_conv postorder.simps(1) preorder.simps(1) |
|
221 |
reflect.simps(1)) |
|
222 |
by (metis append.simps(1) append.simps(2) postorder.simps(2) preorder.simps(2) |
|
223 |
reflect.simps(2) rev.simps(2) rev_append rev_swap) |
|
23449 | 224 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
225 |
declare [[ atp_problem_prefix = "BT__inorder_reflect" ]] |
36484 | 226 |
|
23449 | 227 |
lemma inorder_reflect: "inorder (reflect t) = rev (inorder t)" |
36484 | 228 |
apply (induct t) |
229 |
apply (metis Nil_is_rev_conv inorder.simps(1) reflect.simps(1)) |
|
230 |
by simp |
|
231 |
(* Slow: |
|
232 |
by (metis append.simps(1) append_eq_append_conv2 inorder.simps(2) |
|
233 |
reflect.simps(2) rev.simps(2) rev_append) |
|
234 |
*) |
|
23449 | 235 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
236 |
declare [[ atp_problem_prefix = "BT__postorder_reflect" ]] |
36484 | 237 |
|
23449 | 238 |
lemma postorder_reflect: "postorder (reflect t) = rev (preorder t)" |
36484 | 239 |
apply (induct t) |
240 |
apply (metis Nil_is_rev_conv postorder.simps(1) preorder.simps(1) |
|
241 |
reflect.simps(1)) |
|
242 |
by (metis preorder_reflect reflect_reflect_ident rev_swap) |
|
23449 | 243 |
|
244 |
text {* |
|
36484 | 245 |
Analogues of the standard properties of the append function for lists. |
23449 | 246 |
*} |
247 |
||
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
248 |
declare [[ atp_problem_prefix = "BT__appnd_assoc" ]] |
36484 | 249 |
|
250 |
lemma appnd_assoc [simp]: "appnd (appnd t1 t2) t3 = appnd t1 (appnd t2 t3)" |
|
251 |
apply (induct t1) |
|
252 |
apply (metis appnd.simps(1)) |
|
253 |
by (metis appnd.simps(2)) |
|
23449 | 254 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
255 |
declare [[ atp_problem_prefix = "BT__appnd_Lf2" ]] |
36484 | 256 |
|
23449 | 257 |
lemma appnd_Lf2 [simp]: "appnd t Lf = t" |
36484 | 258 |
apply (induct t) |
259 |
apply (metis appnd.simps(1)) |
|
260 |
by (metis appnd.simps(2)) |
|
261 |
||
262 |
declare max_add_distrib_left [simp] |
|
23449 | 263 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
264 |
declare [[ atp_problem_prefix = "BT__depth_appnd" ]] |
36484 | 265 |
|
23449 | 266 |
lemma depth_appnd [simp]: "depth (appnd t1 t2) = depth t1 + depth t2" |
36484 | 267 |
apply (induct t1) |
268 |
apply (metis appnd.simps(1) depth.simps(1) plus_nat.simps(1)) |
|
269 |
by simp |
|
23449 | 270 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
271 |
declare [[ atp_problem_prefix = "BT__n_leaves_appnd" ]] |
36484 | 272 |
|
23449 | 273 |
lemma n_leaves_appnd [simp]: |
274 |
"n_leaves (appnd t1 t2) = n_leaves t1 * n_leaves t2" |
|
36484 | 275 |
apply (induct t1) |
276 |
apply (metis appnd.simps(1) n_leaves.simps(1) nat_mult_1 plus_nat.simps(1) |
|
277 |
semiring_norm(111)) |
|
278 |
by (simp add: left_distrib) |
|
23449 | 279 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
31790
diff
changeset
|
280 |
declare [[ atp_problem_prefix = "BT__bt_map_appnd" ]] |
36484 | 281 |
|
26312 | 282 |
lemma (*bt_map_appnd:*) |
23449 | 283 |
"bt_map f (appnd t1 t2) = appnd (bt_map f t1) (bt_map f t2)" |
36484 | 284 |
apply (induct t1) |
285 |
apply (metis appnd.simps(1) bt_map.simps(1)) |
|
286 |
by (metis bt_map_appnd) |
|
23449 | 287 |
|
288 |
end |