author | wenzelm |
Mon, 18 Nov 2013 19:56:34 +0100 | |
changeset 54511 | 1fd24c96ce9b |
parent 46900 | 73555abfa267 |
child 58871 | c399ae4b836f |
permissions | -rw-r--r-- |
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
1 |
(* Title: ZF/Induct/Binary_Trees.thy |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
3 |
Copyright 1992 University of Cambridge |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
4 |
*) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
5 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
6 |
header {* Binary trees *} |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
7 |
|
16417 | 8 |
theory Binary_Trees imports Main begin |
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
9 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
10 |
subsection {* Datatype definition *} |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
11 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
12 |
consts |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
13 |
bt :: "i => i" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
14 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
15 |
datatype "bt(A)" = |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
16 |
Lf | Br ("a \<in> A", "t1 \<in> bt(A)", "t2 \<in> bt(A)") |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
17 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
18 |
declare bt.intros [simp] |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
19 |
|
18415 | 20 |
lemma Br_neq_left: "l \<in> bt(A) ==> Br(x, l, r) \<noteq> l" |
20503 | 21 |
by (induct arbitrary: x r set: bt) auto |
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
22 |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
35762
diff
changeset
|
23 |
lemma Br_iff: "Br(a, l, r) = Br(a', l', r') \<longleftrightarrow> a = a' & l = l' & r = r'" |
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
24 |
-- "Proving a freeness theorem." |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
25 |
by (fast elim!: bt.free_elims) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
26 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
27 |
inductive_cases BrE: "Br(a, l, r) \<in> bt(A)" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
28 |
-- "An elimination rule, for type-checking." |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
29 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
30 |
text {* |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
31 |
\medskip Lemmas to justify using @{term bt} in other recursive type |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
32 |
definitions. |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
33 |
*} |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
34 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
35 |
lemma bt_mono: "A \<subseteq> B ==> bt(A) \<subseteq> bt(B)" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
36 |
apply (unfold bt.defs) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
37 |
apply (rule lfp_mono) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
38 |
apply (rule bt.bnd_mono)+ |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
39 |
apply (rule univ_mono basic_monos | assumption)+ |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
40 |
done |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
41 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
42 |
lemma bt_univ: "bt(univ(A)) \<subseteq> univ(A)" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
43 |
apply (unfold bt.defs bt.con_defs) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
44 |
apply (rule lfp_lowerbound) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
45 |
apply (rule_tac [2] A_subset_univ [THEN univ_mono]) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
46 |
apply (fast intro!: zero_in_univ Inl_in_univ Inr_in_univ Pair_in_univ) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
47 |
done |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
48 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
49 |
lemma bt_subset_univ: "A \<subseteq> univ(B) ==> bt(A) \<subseteq> univ(B)" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
50 |
apply (rule subset_trans) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
51 |
apply (erule bt_mono) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
52 |
apply (rule bt_univ) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
53 |
done |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
54 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
55 |
lemma bt_rec_type: |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
56 |
"[| t \<in> bt(A); |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
57 |
c \<in> C(Lf); |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
58 |
!!x y z r s. [| x \<in> A; y \<in> bt(A); z \<in> bt(A); r \<in> C(y); s \<in> C(z) |] ==> |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
59 |
h(x, y, z, r, s) \<in> C(Br(x, y, z)) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
60 |
|] ==> bt_rec(c, h, t) \<in> C(t)" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
61 |
-- {* Type checking for recursor -- example only; not really needed. *} |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
62 |
apply (induct_tac t) |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
63 |
apply simp_all |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
64 |
done |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
65 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
66 |
|
14157 | 67 |
subsection {* Number of nodes, with an example of tail-recursion *} |
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
68 |
|
14157 | 69 |
consts n_nodes :: "i => i" |
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
70 |
primrec |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
71 |
"n_nodes(Lf) = 0" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
72 |
"n_nodes(Br(a, l, r)) = succ(n_nodes(l) #+ n_nodes(r))" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
73 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
74 |
lemma n_nodes_type [simp]: "t \<in> bt(A) ==> n_nodes(t) \<in> nat" |
18415 | 75 |
by (induct set: bt) auto |
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
76 |
|
14157 | 77 |
consts n_nodes_aux :: "i => i" |
78 |
primrec |
|
79 |
"n_nodes_aux(Lf) = (\<lambda>k \<in> nat. k)" |
|
18415 | 80 |
"n_nodes_aux(Br(a, l, r)) = |
14157 | 81 |
(\<lambda>k \<in> nat. n_nodes_aux(r) ` (n_nodes_aux(l) ` succ(k)))" |
82 |
||
18415 | 83 |
lemma n_nodes_aux_eq: |
84 |
"t \<in> bt(A) ==> k \<in> nat ==> n_nodes_aux(t)`k = n_nodes(t) #+ k" |
|
20503 | 85 |
apply (induct arbitrary: k set: bt) |
18415 | 86 |
apply simp |
87 |
apply (atomize, simp) |
|
88 |
done |
|
14157 | 89 |
|
24893 | 90 |
definition |
91 |
n_nodes_tail :: "i => i" where |
|
18415 | 92 |
"n_nodes_tail(t) == n_nodes_aux(t) ` 0" |
14157 | 93 |
|
94 |
lemma "t \<in> bt(A) ==> n_nodes_tail(t) = n_nodes(t)" |
|
18415 | 95 |
by (simp add: n_nodes_tail_def n_nodes_aux_eq) |
14157 | 96 |
|
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
97 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
98 |
subsection {* Number of leaves *} |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
99 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
100 |
consts |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
101 |
n_leaves :: "i => i" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
102 |
primrec |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
103 |
"n_leaves(Lf) = 1" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
104 |
"n_leaves(Br(a, l, r)) = n_leaves(l) #+ n_leaves(r)" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
105 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
106 |
lemma n_leaves_type [simp]: "t \<in> bt(A) ==> n_leaves(t) \<in> nat" |
18415 | 107 |
by (induct set: bt) auto |
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
108 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
109 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
110 |
subsection {* Reflecting trees *} |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
111 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
112 |
consts |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
113 |
bt_reflect :: "i => i" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
114 |
primrec |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
115 |
"bt_reflect(Lf) = Lf" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
116 |
"bt_reflect(Br(a, l, r)) = Br(a, bt_reflect(r), bt_reflect(l))" |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
117 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
118 |
lemma bt_reflect_type [simp]: "t \<in> bt(A) ==> bt_reflect(t) \<in> bt(A)" |
18415 | 119 |
by (induct set: bt) auto |
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
120 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
121 |
text {* |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
122 |
\medskip Theorems about @{term n_leaves}. |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
123 |
*} |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
124 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
125 |
lemma n_leaves_reflect: "t \<in> bt(A) ==> n_leaves(bt_reflect(t)) = n_leaves(t)" |
46900 | 126 |
by (induct set: bt) (simp_all add: add_commute) |
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
127 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
128 |
lemma n_leaves_nodes: "t \<in> bt(A) ==> n_leaves(t) = succ(n_nodes(t))" |
46900 | 129 |
by (induct set: bt) simp_all |
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
130 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
131 |
text {* |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
132 |
Theorems about @{term bt_reflect}. |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
133 |
*} |
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
134 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
135 |
lemma bt_reflect_bt_reflect_ident: "t \<in> bt(A) ==> bt_reflect(bt_reflect(t)) = t" |
18415 | 136 |
by (induct set: bt) simp_all |
12194
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
137 |
|
13909cb72129
converted datatype examples moved from ZF/ex to ZF/Induct;
wenzelm
parents:
diff
changeset
|
138 |
end |