| author | wenzelm | 
| Thu, 13 Apr 2000 17:49:42 +0200 | |
| changeset 8711 | 00ec2ba9174d | 
| parent 6144 | 7d38744313c8 | 
| child 11316 | b4e71bd751e4 | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/ex/BT.ML | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 515 | 4 | Copyright 1994 University of Cambridge | 
| 0 | 5 | |
| 6 | Datatype definition of binary trees | |
| 7 | *) | |
| 8 | ||
| 6046 | 9 | Addsimps bt.intrs; | 
| 10 | ||
| 6112 | 11 | Goal "l : bt(A) ==> ALL x r. Br(x,l,r) ~= l"; | 
| 12 | by (induct_tac "l" 1); | |
| 6144 | 13 | by Auto_tac; | 
| 6112 | 14 | qed_spec_mp "Br_neq_left"; | 
| 15 | ||
| 16 | (*Proving a freeness theorem*) | |
| 17 | val Br_iff = bt.mk_free "Br(a,l,r) = Br(a',l',r') <-> a=a' & l=l' & r=r'"; | |
| 18 | ||
| 19 | (*An elimination rule, for type-checking*) | |
| 6141 | 20 | val BrE = bt.mk_cases "Br(a,l,r) : bt(A)"; | 
| 6112 | 21 | |
| 0 | 22 | (** Lemmas to justify using "bt" in other recursive type definitions **) | 
| 23 | ||
| 5137 | 24 | Goalw bt.defs "A<=B ==> bt(A) <= bt(B)"; | 
| 0 | 25 | by (rtac lfp_mono 1); | 
| 515 | 26 | by (REPEAT (rtac bt.bnd_mono 1)); | 
| 0 | 27 | by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 28 | qed "bt_mono"; | 
| 0 | 29 | |
| 5068 | 30 | Goalw (bt.defs@bt.con_defs) "bt(univ(A)) <= univ(A)"; | 
| 0 | 31 | by (rtac lfp_lowerbound 1); | 
| 32 | by (rtac (A_subset_univ RS univ_mono) 2); | |
| 4091 | 33 | by (fast_tac (claset() addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ, | 
| 1461 | 34 | Pair_in_univ]) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 35 | qed "bt_univ"; | 
| 0 | 36 | |
| 6046 | 37 | bind_thm ("bt_subset_univ", [bt_mono, bt_univ] MRS subset_trans);
 | 
| 515 | 38 | |
| 2469 | 39 | |
| 6112 | 40 | (*Type checking for recursor -- example only; not really needed*) | 
| 6065 | 41 | val major::prems = goal BT.thy | 
| 515 | 42 | "[| t: bt(A); \ | 
| 43 | \ c: C(Lf); \ | |
| 44 | \ !!x y z r s. [| x:A; y:bt(A); z:bt(A); r:C(y); s:C(z) |] ==> \ | |
| 1461 | 45 | \ h(x,y,z,r,s): C(Br(x,y,z)) \ | 
| 6046 | 46 | \ |] ==> bt_rec(c,h,t) : C(t)"; | 
| 6065 | 47 | (*instead of induct_tac "t", since t: bt(A) isn't an assumption*) | 
| 48 | by (rtac (major RS bt.induct) 1); | |
| 4091 | 49 | by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); | 
| 760 | 50 | qed "bt_rec_type"; | 
| 515 | 51 | |
| 52 | (** n_nodes **) | |
| 53 | ||
| 6046 | 54 | Goal "t: bt(A) ==> n_nodes(t) : nat"; | 
| 6065 | 55 | by (induct_tac "t" 1); | 
| 6046 | 56 | by Auto_tac; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 57 | qed "n_nodes_type"; | 
| 515 | 58 | |
| 59 | ||
| 60 | (** n_leaves **) | |
| 61 | ||
| 6046 | 62 | Goal "t: bt(A) ==> n_leaves(t) : nat"; | 
| 6065 | 63 | by (induct_tac "t" 1); | 
| 6046 | 64 | by Auto_tac; | 
| 760 | 65 | qed "n_leaves_type"; | 
| 515 | 66 | |
| 67 | (** bt_reflect **) | |
| 68 | ||
| 6046 | 69 | Goal "t: bt(A) ==> bt_reflect(t) : bt(A)"; | 
| 6065 | 70 | by (induct_tac "t" 1); | 
| 6046 | 71 | by Auto_tac; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 72 | qed "bt_reflect_type"; | 
| 515 | 73 | |
| 74 | ||
| 75 | (** BT simplification **) | |
| 76 | ||
| 77 | ||
| 6112 | 78 | Addsimps [n_nodes_type, n_leaves_type, bt_reflect_type]; | 
| 515 | 79 | |
| 80 | ||
| 81 | (*** theorems about n_leaves ***) | |
| 82 | ||
| 6065 | 83 | Goal "t: bt(A) ==> n_leaves(bt_reflect(t)) = n_leaves(t)"; | 
| 84 | by (induct_tac "t" 1); | |
| 4091 | 85 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_commute, n_leaves_type]))); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 86 | qed "n_leaves_reflect"; | 
| 515 | 87 | |
| 6065 | 88 | Goal "t: bt(A) ==> n_leaves(t) = succ(n_nodes(t))"; | 
| 89 | by (induct_tac "t" 1); | |
| 4091 | 90 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_succ_right]))); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 91 | qed "n_leaves_nodes"; | 
| 515 | 92 | |
| 93 | (*** theorems about bt_reflect ***) | |
| 94 | ||
| 6065 | 95 | Goal "t: bt(A) ==> bt_reflect(bt_reflect(t))=t"; | 
| 96 | by (induct_tac "t" 1); | |
| 2469 | 97 | by (ALLGOALS Asm_simp_tac); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 98 | qed "bt_reflect_bt_reflect_ident"; | 
| 515 | 99 | |
| 100 |