| author | haftmann | 
| Tue, 17 Aug 2010 14:33:44 +0200 | |
| changeset 38463 | 0e4565062017 | 
| parent 35762 | af3ff2ba4c54 | 
| child 41526 | 54b4686704af | 
| permissions | -rw-r--r-- | 
| 12201 | 1 | (* Title: ZF/Induct/Tree_Forest.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 1994 University of Cambridge | |
| 4 | *) | |
| 5 | ||
| 6 | header {* Trees and forests, a mutually recursive type definition *}
 | |
| 7 | ||
| 16417 | 8 | theory Tree_Forest imports Main begin | 
| 12201 | 9 | |
| 10 | subsection {* Datatype definition *}
 | |
| 11 | ||
| 12 | consts | |
| 13 | tree :: "i => i" | |
| 14 | forest :: "i => i" | |
| 15 | tree_forest :: "i => i" | |
| 16 | ||
| 17 | datatype "tree(A)" = Tcons ("a \<in> A", "f \<in> forest(A)")
 | |
| 18 |   and "forest(A)" = Fnil | Fcons ("t \<in> tree(A)", "f \<in> forest(A)")
 | |
| 19 | ||
| 18460 | 20 | (* FIXME *) | 
| 21 | lemmas tree'induct = | |
| 22 | tree_forest.mutual_induct [THEN conjunct1, THEN spec, THEN [2] rev_mp, of concl: _ t, standard, consumes 1] | |
| 23 | and forest'induct = | |
| 24 | tree_forest.mutual_induct [THEN conjunct2, THEN spec, THEN [2] rev_mp, of concl: _ f, standard, consumes 1] | |
| 25 | ||
| 12201 | 26 | declare tree_forest.intros [simp, TC] | 
| 27 | ||
| 12216 | 28 | lemma tree_def: "tree(A) == Part(tree_forest(A), Inl)" | 
| 12201 | 29 | by (simp only: tree_forest.defs) | 
| 30 | ||
| 12216 | 31 | lemma forest_def: "forest(A) == Part(tree_forest(A), Inr)" | 
| 12201 | 32 | by (simp only: tree_forest.defs) | 
| 33 | ||
| 34 | ||
| 35 | text {*
 | |
| 12205 | 36 |   \medskip @{term "tree_forest(A)"} as the union of @{term "tree(A)"}
 | 
| 12201 | 37 |   and @{term "forest(A)"}.
 | 
| 38 | *} | |
| 39 | ||
| 40 | lemma tree_subset_TF: "tree(A) \<subseteq> tree_forest(A)" | |
| 41 | apply (unfold tree_forest.defs) | |
| 42 | apply (rule Part_subset) | |
| 43 | done | |
| 44 | ||
| 12243 | 45 | lemma treeI [TC]: "x \<in> tree(A) ==> x \<in> tree_forest(A)" | 
| 12201 | 46 | by (rule tree_subset_TF [THEN subsetD]) | 
| 47 | ||
| 48 | lemma forest_subset_TF: "forest(A) \<subseteq> tree_forest(A)" | |
| 49 | apply (unfold tree_forest.defs) | |
| 50 | apply (rule Part_subset) | |
| 51 | done | |
| 52 | ||
| 13535 | 53 | lemma treeI' [TC]: "x \<in> forest(A) ==> x \<in> tree_forest(A)" | 
| 12201 | 54 | by (rule forest_subset_TF [THEN subsetD]) | 
| 55 | ||
| 56 | lemma TF_equals_Un: "tree(A) \<union> forest(A) = tree_forest(A)" | |
| 57 | apply (insert tree_subset_TF forest_subset_TF) | |
| 58 | apply (auto intro!: equalityI tree_forest.intros elim: tree_forest.cases) | |
| 59 | done | |
| 60 | ||
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12610diff
changeset | 61 | lemma | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12610diff
changeset | 62 | notes rews = tree_forest.con_defs tree_def forest_def | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12610diff
changeset | 63 | shows | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12610diff
changeset | 64 | tree_forest_unfold: "tree_forest(A) = | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12610diff
changeset | 65 |       (A \<times> forest(A)) + ({0} + tree(A) \<times> forest(A))"
 | 
| 12201 | 66 |     -- {* NOT useful, but interesting \dots *}
 | 
| 67 | apply (unfold tree_def forest_def) | |
| 68 | apply (fast intro!: tree_forest.intros [unfolded rews, THEN PartD1] | |
| 69 | elim: tree_forest.cases [unfolded rews]) | |
| 70 | done | |
| 71 | ||
| 72 | lemma tree_forest_unfold': | |
| 73 | "tree_forest(A) = | |
| 74 | A \<times> Part(tree_forest(A), \<lambda>w. Inr(w)) + | |
| 75 |     {0} + Part(tree_forest(A), \<lambda>w. Inl(w)) * Part(tree_forest(A), \<lambda>w. Inr(w))"
 | |
| 76 | by (rule tree_forest_unfold [unfolded tree_def forest_def]) | |
| 77 | ||
| 78 | lemma tree_unfold: "tree(A) = {Inl(x). x \<in> A \<times> forest(A)}"
 | |
| 79 | apply (unfold tree_def forest_def) | |
| 80 | apply (rule Part_Inl [THEN subst]) | |
| 81 | apply (rule tree_forest_unfold' [THEN subst_context]) | |
| 82 | done | |
| 83 | ||
| 84 | lemma forest_unfold: "forest(A) = {Inr(x). x \<in> {0} + tree(A)*forest(A)}"
 | |
| 85 | apply (unfold tree_def forest_def) | |
| 86 | apply (rule Part_Inr [THEN subst]) | |
| 87 | apply (rule tree_forest_unfold' [THEN subst_context]) | |
| 88 | done | |
| 89 | ||
| 90 | text {*
 | |
| 91 | \medskip Type checking for recursor: Not needed; possibly interesting? | |
| 92 | *} | |
| 93 | ||
| 94 | lemma TF_rec_type: | |
| 95 | "[| z \<in> tree_forest(A); | |
| 96 | !!x f r. [| x \<in> A; f \<in> forest(A); r \<in> C(f) | |
| 12243 | 97 | |] ==> b(x,f,r) \<in> C(Tcons(x,f)); | 
| 12201 | 98 | c \<in> C(Fnil); | 
| 12243 | 99 | !!t f r1 r2. [| t \<in> tree(A); f \<in> forest(A); r1 \<in> C(t); r2 \<in> C(f) | 
| 100 | |] ==> d(t,f,r1,r2) \<in> C(Fcons(t,f)) | |
| 12201 | 101 | |] ==> tree_forest_rec(b,c,d,z) \<in> C(z)" | 
| 102 | by (induct_tac z) simp_all | |
| 103 | ||
| 104 | lemma tree_forest_rec_type: | |
| 105 | "[| !!x f r. [| x \<in> A; f \<in> forest(A); r \<in> D(f) | |
| 12243 | 106 | |] ==> b(x,f,r) \<in> C(Tcons(x,f)); | 
| 12201 | 107 | c \<in> D(Fnil); | 
| 12243 | 108 | !!t f r1 r2. [| t \<in> tree(A); f \<in> forest(A); r1 \<in> C(t); r2 \<in> D(f) | 
| 109 | |] ==> d(t,f,r1,r2) \<in> D(Fcons(t,f)) | |
| 110 | |] ==> (\<forall>t \<in> tree(A). tree_forest_rec(b,c,d,t) \<in> C(t)) \<and> | |
| 12201 | 111 | (\<forall>f \<in> forest(A). tree_forest_rec(b,c,d,f) \<in> D(f))" | 
| 112 |     -- {* Mutually recursive version. *}
 | |
| 12243 | 113 | apply (unfold Ball_def) | 
| 12201 | 114 | apply (rule tree_forest.mutual_induct) | 
| 115 | apply simp_all | |
| 116 | done | |
| 117 | ||
| 118 | ||
| 119 | subsection {* Operations *}
 | |
| 120 | ||
| 121 | consts | |
| 122 | map :: "[i => i, i] => i" | |
| 123 | size :: "i => i" | |
| 124 | preorder :: "i => i" | |
| 125 | list_of_TF :: "i => i" | |
| 126 | of_list :: "i => i" | |
| 127 | reflect :: "i => i" | |
| 128 | ||
| 129 | primrec | |
| 130 | "list_of_TF (Tcons(x,f)) = [Tcons(x,f)]" | |
| 131 | "list_of_TF (Fnil) = []" | |
| 132 | "list_of_TF (Fcons(t,tf)) = Cons (t, list_of_TF(tf))" | |
| 133 | ||
| 134 | primrec | |
| 135 | "of_list([]) = Fnil" | |
| 136 | "of_list(Cons(t,l)) = Fcons(t, of_list(l))" | |
| 137 | ||
| 138 | primrec | |
| 139 | "map (h, Tcons(x,f)) = Tcons(h(x), map(h,f))" | |
| 140 | "map (h, Fnil) = Fnil" | |
| 141 | "map (h, Fcons(t,tf)) = Fcons (map(h, t), map(h, tf))" | |
| 142 | ||
| 143 | primrec | |
| 144 | "size (Tcons(x,f)) = succ(size(f))" | |
| 145 | "size (Fnil) = 0" | |
| 146 | "size (Fcons(t,tf)) = size(t) #+ size(tf)" | |
| 147 | ||
| 148 | primrec | |
| 149 | "preorder (Tcons(x,f)) = Cons(x, preorder(f))" | |
| 150 | "preorder (Fnil) = Nil" | |
| 151 | "preorder (Fcons(t,tf)) = preorder(t) @ preorder(tf)" | |
| 152 | ||
| 153 | primrec | |
| 154 | "reflect (Tcons(x,f)) = Tcons(x, reflect(f))" | |
| 155 | "reflect (Fnil) = Fnil" | |
| 156 | "reflect (Fcons(t,tf)) = | |
| 157 | of_list (list_of_TF (reflect(tf)) @ Cons(reflect(t), Nil))" | |
| 158 | ||
| 159 | ||
| 160 | text {*
 | |
| 161 |   \medskip @{text list_of_TF} and @{text of_list}.
 | |
| 162 | *} | |
| 163 | ||
| 164 | lemma list_of_TF_type [TC]: | |
| 165 | "z \<in> tree_forest(A) ==> list_of_TF(z) \<in> list(tree(A))" | |
| 18460 | 166 | by (induct set: tree_forest) simp_all | 
| 12201 | 167 | |
| 12243 | 168 | lemma of_list_type [TC]: "l \<in> list(tree(A)) ==> of_list(l) \<in> forest(A)" | 
| 18460 | 169 | by (induct set: list) simp_all | 
| 12201 | 170 | |
| 171 | text {*
 | |
| 172 |   \medskip @{text map}.
 | |
| 173 | *} | |
| 174 | ||
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12610diff
changeset | 175 | lemma | 
| 18460 | 176 | assumes "!!x. x \<in> A ==> h(x): B" | 
| 12937 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12610diff
changeset | 177 | shows map_tree_type: "t \<in> tree(A) ==> map(h,t) \<in> tree(B)" | 
| 
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
 wenzelm parents: 
12610diff
changeset | 178 | and map_forest_type: "f \<in> forest(A) ==> map(h,f) \<in> forest(B)" | 
| 18460 | 179 | using prems | 
| 180 | by (induct rule: tree'induct forest'induct) simp_all | |
| 12201 | 181 | |
| 182 | text {*
 | |
| 183 |   \medskip @{text size}.
 | |
| 184 | *} | |
| 185 | ||
| 186 | lemma size_type [TC]: "z \<in> tree_forest(A) ==> size(z) \<in> nat" | |
| 18460 | 187 | by (induct set: tree_forest) simp_all | 
| 12201 | 188 | |
| 189 | ||
| 190 | text {*
 | |
| 191 |   \medskip @{text preorder}.
 | |
| 192 | *} | |
| 193 | ||
| 194 | lemma preorder_type [TC]: "z \<in> tree_forest(A) ==> preorder(z) \<in> list(A)" | |
| 18460 | 195 | by (induct set: tree_forest) simp_all | 
| 12201 | 196 | |
| 197 | ||
| 198 | text {*
 | |
| 199 |   \medskip Theorems about @{text list_of_TF} and @{text of_list}.
 | |
| 200 | *} | |
| 201 | ||
| 18415 | 202 | lemma forest_induct [consumes 1, case_names Fnil Fcons]: | 
| 12201 | 203 | "[| f \<in> forest(A); | 
| 204 | R(Fnil); | |
| 205 | !!t f. [| t \<in> tree(A); f \<in> forest(A); R(f) |] ==> R(Fcons(t,f)) | |
| 206 | |] ==> R(f)" | |
| 207 |   -- {* Essentially the same as list induction. *}
 | |
| 12610 | 208 | apply (erule tree_forest.mutual_induct | 
| 209 | [THEN conjunct2, THEN spec, THEN [2] rev_mp]) | |
| 12201 | 210 | apply (rule TrueI) | 
| 211 | apply simp | |
| 212 | apply simp | |
| 213 | done | |
| 214 | ||
| 215 | lemma forest_iso: "f \<in> forest(A) ==> of_list(list_of_TF(f)) = f" | |
| 18415 | 216 | by (induct rule: forest_induct) simp_all | 
| 12201 | 217 | |
| 218 | lemma tree_list_iso: "ts: list(tree(A)) ==> list_of_TF(of_list(ts)) = ts" | |
| 18415 | 219 | by (induct set: list) simp_all | 
| 12201 | 220 | |
| 221 | ||
| 222 | text {*
 | |
| 223 |   \medskip Theorems about @{text map}.
 | |
| 224 | *} | |
| 225 | ||
| 226 | lemma map_ident: "z \<in> tree_forest(A) ==> map(\<lambda>u. u, z) = z" | |
| 18460 | 227 | by (induct set: tree_forest) simp_all | 
| 12201 | 228 | |
| 12216 | 229 | lemma map_compose: | 
| 230 | "z \<in> tree_forest(A) ==> map(h, map(j,z)) = map(\<lambda>u. h(j(u)), z)" | |
| 18460 | 231 | by (induct set: tree_forest) simp_all | 
| 12201 | 232 | |
| 233 | ||
| 234 | text {*
 | |
| 235 |   \medskip Theorems about @{text size}.
 | |
| 236 | *} | |
| 237 | ||
| 238 | lemma size_map: "z \<in> tree_forest(A) ==> size(map(h,z)) = size(z)" | |
| 18460 | 239 | by (induct set: tree_forest) simp_all | 
| 12201 | 240 | |
| 241 | lemma size_length: "z \<in> tree_forest(A) ==> size(z) = length(preorder(z))" | |
| 18460 | 242 | by (induct set: tree_forest) (simp_all add: length_app) | 
| 12201 | 243 | |
| 244 | text {*
 | |
| 245 |   \medskip Theorems about @{text preorder}.
 | |
| 246 | *} | |
| 247 | ||
| 248 | lemma preorder_map: | |
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: 
18460diff
changeset | 249 | "z \<in> tree_forest(A) ==> preorder(map(h,z)) = List_ZF.map(h, preorder(z))" | 
| 18460 | 250 | by (induct set: tree_forest) (simp_all add: map_app_distrib) | 
| 12201 | 251 | |
| 252 | end |