| author | wenzelm | 
| Thu, 13 Apr 2000 17:49:42 +0200 | |
| changeset 8711 | 00ec2ba9174d | 
| parent 6154 | 6a00a5baef2b | 
| child 11316 | b4e71bd751e4 | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/ex/tf.ML | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 6 | Trees & forests, a mutually recursive type definition. | |
| 7 | *) | |
| 8 | ||
| 6046 | 9 | Addsimps tree_forest.intrs; | 
| 6153 | 10 | AddTCs tree_forest.intrs; | 
| 0 | 11 | |
| 12 | (** tree_forest(A) as the union of tree(A) and forest(A) **) | |
| 13 | ||
| 739 | 14 | val [_, tree_def, forest_def] = tree_forest.defs; | 
| 15 | ||
| 5068 | 16 | Goalw [tree_def] "tree(A) <= tree_forest(A)"; | 
| 0 | 17 | by (rtac Part_subset 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 18 | qed "tree_subset_TF"; | 
| 0 | 19 | |
| 5068 | 20 | Goalw [forest_def] "forest(A) <= tree_forest(A)"; | 
| 0 | 21 | by (rtac Part_subset 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 22 | qed "forest_subset_TF"; | 
| 0 | 23 | |
| 5068 | 24 | Goal "tree(A) Un forest(A) = tree_forest(A)"; | 
| 739 | 25 | by (safe_tac (subset_cs addSIs [equalityI, tree_subset_TF, forest_subset_TF])); | 
| 4091 | 26 | by (fast_tac (claset() addSIs tree_forest.intrs addEs [tree_forest.elim]) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 27 | qed "TF_equals_Un"; | 
| 0 | 28 | |
| 29 | (** NOT useful, but interesting... **) | |
| 30 | ||
| 5068 | 31 | Goalw [tree_def, forest_def] | 
| 434 | 32 |     "tree_forest(A) = (A*forest(A)) + ({0} + tree(A)*forest(A))";
 | 
| 529 
f0d16216e394
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 lcp parents: 
515diff
changeset | 33 | let open tree_forest; | 
| 
f0d16216e394
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 lcp parents: 
515diff
changeset | 34 | val rew = rewrite_rule (con_defs @ tl defs) in | 
| 4091 | 35 | by (fast_tac (claset() addSIs (map rew intrs RL [PartD1]) addEs [rew elim]) 1) | 
| 529 
f0d16216e394
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 lcp parents: 
515diff
changeset | 36 | end; | 
| 760 | 37 | qed "tree_forest_unfold"; | 
| 434 | 38 | |
| 739 | 39 | val tree_forest_unfold' = rewrite_rule [tree_def, forest_def] | 
| 529 
f0d16216e394
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 lcp parents: 
515diff
changeset | 40 | tree_forest_unfold; | 
| 434 | 41 | |
| 5068 | 42 | Goalw [tree_def, forest_def] | 
| 515 | 43 |     "tree(A) = {Inl(x). x: A*forest(A)}";
 | 
| 438 | 44 | by (rtac (Part_Inl RS subst) 1); | 
| 45 | by (rtac (tree_forest_unfold' RS subst_context) 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 46 | qed "tree_unfold"; | 
| 0 | 47 | |
| 5068 | 48 | Goalw [tree_def, forest_def] | 
| 515 | 49 |     "forest(A) = {Inr(x). x: {0} + tree(A)*forest(A)}";
 | 
| 438 | 50 | by (rtac (Part_Inr RS subst) 1); | 
| 51 | by (rtac (tree_forest_unfold' RS subst_context) 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 52 | qed "forest_unfold"; | 
| 0 | 53 | |
| 434 | 54 | |
| 6046 | 55 | (** Type checking for recursor: Not needed; possibly interesting (??) **) | 
| 515 | 56 | |
| 57 | val major::prems = goal TF.thy | |
| 58 | "[| z: tree_forest(A); \ | |
| 1461 | 59 | \ !!x f r. [| x: A; f: forest(A); r: C(f) \ | 
| 60 | \ |] ==> b(x,f,r): C(Tcons(x,f)); \ | |
| 61 | \ c : C(Fnil); \ | |
| 515 | 62 | \ !!t f r1 r2. [| t: tree(A); f: forest(A); r1: C(t); r2: C(f) \ | 
| 1461 | 63 | \ |] ==> d(t,f,r1,r2): C(Fcons(t,f)) \ | 
| 6046 | 64 | \ |] ==> tree_forest_rec(b,c,d,z) : C(z)"; | 
| 515 | 65 | by (rtac (major RS tree_forest.induct) 1); | 
| 4091 | 66 | by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); | 
| 760 | 67 | qed "TF_rec_type"; | 
| 515 | 68 | |
| 69 | (*Mutually recursive version*) | |
| 70 | val prems = goal TF.thy | |
| 1461 | 71 | "[| !!x f r. [| x: A; f: forest(A); r: D(f) \ | 
| 72 | \ |] ==> b(x,f,r): C(Tcons(x,f)); \ | |
| 73 | \ c : D(Fnil); \ | |
| 515 | 74 | \ !!t f r1 r2. [| t: tree(A); f: forest(A); r1: C(t); r2: D(f) \ | 
| 1461 | 75 | \ |] ==> d(t,f,r1,r2): D(Fcons(t,f)) \ | 
| 6046 | 76 | \ |] ==> (ALL t:tree(A). tree_forest_rec(b,c,d,t) : C(t)) & \ | 
| 77 | \ (ALL f: forest(A). tree_forest_rec(b,c,d,f) : D(f))"; | |
| 515 | 78 | by (rewtac Ball_def); | 
| 79 | by (rtac tree_forest.mutual_induct 1); | |
| 4091 | 80 | by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 81 | qed "tree_forest_rec_type"; | 
| 515 | 82 | |
| 83 | ||
| 6046 | 84 | (** list_of_TF and of_list **) | 
| 515 | 85 | |
| 6046 | 86 | Goal "z: tree_forest(A) ==> list_of_TF(z) : list(tree(A))"; | 
| 87 | by (etac tree_forest.induct 1); | |
| 88 | by (ALLGOALS Asm_simp_tac); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 89 | qed "list_of_TF_type"; | 
| 515 | 90 | |
| 6046 | 91 | Goal "l: list(tree(A)) ==> of_list(l) : forest(A)"; | 
| 92 | by (etac list.induct 1); | |
| 93 | by (ALLGOALS Asm_simp_tac); | |
| 94 | qed "of_list_type"; | |
| 515 | 95 | |
| 96 | ||
| 6046 | 97 | (** map **) | 
| 515 | 98 | |
| 6046 | 99 | val prems = Goal | 
| 515 | 100 | "[| !!x. x: A ==> h(x): B |] ==> \ | 
| 6046 | 101 | \ (ALL t:tree(A). map(h,t) : tree(B)) & \ | 
| 102 | \ (ALL f: forest(A). map(h,f) : forest(B))"; | |
| 103 | by (rewtac Ball_def); | |
| 104 | by (rtac tree_forest.mutual_induct 1); | |
| 105 | by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); | |
| 106 | qed "map_type"; | |
| 515 | 107 | |
| 108 | ||
| 6046 | 109 | (** size **) | 
| 515 | 110 | |
| 6046 | 111 | Goal "z: tree_forest(A) ==> size(z) : nat"; | 
| 112 | by (etac tree_forest.induct 1); | |
| 113 | by (ALLGOALS Asm_simp_tac); | |
| 114 | qed "size_type"; | |
| 515 | 115 | |
| 116 | ||
| 6046 | 117 | (** preorder **) | 
| 515 | 118 | |
| 6046 | 119 | Goal "z: tree_forest(A) ==> preorder(z) : list(A)"; | 
| 120 | by (etac tree_forest.induct 1); | |
| 121 | by (ALLGOALS Asm_simp_tac); | |
| 122 | qed "preorder_type"; | |
| 515 | 123 | |
| 124 | ||
| 125 | (** Term simplification **) | |
| 126 | ||
| 127 | val treeI = tree_subset_TF RS subsetD | |
| 128 | and forestI = forest_subset_TF RS subsetD; | |
| 129 | ||
| 6154 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6153diff
changeset | 130 | AddTCs [treeI, forestI, list_of_TF_type, map_type, size_type, preorder_type]; | 
| 515 | 131 | |
| 6046 | 132 | (** theorems about list_of_TF and of_list **) | 
| 515 | 133 | |
| 134 | (*essentially the same as list induction*) | |
| 6046 | 135 | val major::prems = Goal | 
| 1461 | 136 | "[| f: forest(A); \ | 
| 515 | 137 | \ R(Fnil); \ | 
| 138 | \ !!t f. [| t: tree(A); f: forest(A); R(f) |] ==> R(Fcons(t,f)) \ | |
| 139 | \ |] ==> R(f)"; | |
| 140 | by (rtac (major RS (tree_forest.mutual_induct RS conjunct2 RS spec RSN (2,rev_mp))) 1); | |
| 141 | by (REPEAT (ares_tac (TrueI::prems) 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 142 | qed "forest_induct"; | 
| 515 | 143 | |
| 6046 | 144 | Goal "f: forest(A) ==> of_list(list_of_TF(f)) = f"; | 
| 515 | 145 | by (etac forest_induct 1); | 
| 2469 | 146 | by (ALLGOALS Asm_simp_tac); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 147 | qed "forest_iso"; | 
| 515 | 148 | |
| 6046 | 149 | Goal "ts: list(tree(A)) ==> list_of_TF(of_list(ts)) = ts"; | 
| 515 | 150 | by (etac list.induct 1); | 
| 2469 | 151 | by (ALLGOALS Asm_simp_tac); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 152 | qed "tree_list_iso"; | 
| 515 | 153 | |
| 6046 | 154 | (** theorems about map **) | 
| 515 | 155 | |
| 6046 | 156 | Goal "z: tree_forest(A) ==> map(%u. u, z) = z"; | 
| 515 | 157 | by (etac tree_forest.induct 1); | 
| 2469 | 158 | by (ALLGOALS Asm_simp_tac); | 
| 6046 | 159 | qed "map_ident"; | 
| 515 | 160 | |
| 6046 | 161 | Goal "z: tree_forest(A) ==> map(h, map(j,z)) = map(%u. h(j(u)), z)"; | 
| 162 | by (etac tree_forest.induct 1); | |
| 163 | by (ALLGOALS Asm_simp_tac); | |
| 164 | qed "map_compose"; | |
| 165 | ||
| 166 | (** theorems about size **) | |
| 167 | ||
| 168 | Goal "z: tree_forest(A) ==> size(map(h,z)) = size(z)"; | |
| 515 | 169 | by (etac tree_forest.induct 1); | 
| 2469 | 170 | by (ALLGOALS Asm_simp_tac); | 
| 6046 | 171 | qed "size_map"; | 
| 515 | 172 | |
| 6046 | 173 | Goal "z: tree_forest(A) ==> size(z) = length(preorder(z))"; | 
| 515 | 174 | by (etac tree_forest.induct 1); | 
| 6154 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6153diff
changeset | 175 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [length_app]))); | 
| 6046 | 176 | qed "size_length"; | 
| 515 | 177 | |
| 6046 | 178 | (** theorems about preorder **) | 
| 515 | 179 | |
| 6046 | 180 | Goal "z: tree_forest(A) ==> preorder(TF.map(h,z)) = List.map(h, preorder(z))"; | 
| 515 | 181 | by (etac tree_forest.induct 1); | 
| 6154 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6153diff
changeset | 182 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [map_app_distrib]))); | 
| 6046 | 183 | qed "preorder_map"; |