| author | paulson | 
| Thu, 28 Mar 1996 10:52:59 +0100 | |
| changeset 1623 | 2b8573c1b1c1 | 
| parent 1461 | 6bcb44e4d6e5 | 
| child 2469 | b50b8c0eec01 | 
| permissions | -rw-r--r-- | 
| 1461 | 1  | 
(* Title: ZF/ex/Ntree.ML  | 
| 486 | 2  | 
ID: $Id$  | 
| 1461 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 486 | 4  | 
Copyright 1994 University of Cambridge  | 
5  | 
||
6  | 
Datatype definition n-ary branching trees  | 
|
7  | 
Demonstrates a simple use of function space in a datatype definition  | 
|
8  | 
||
9  | 
Based upon ex/Term.ML  | 
|
10  | 
*)  | 
|
11  | 
||
| 515 | 12  | 
open Ntree;  | 
| 486 | 13  | 
|
| 539 | 14  | 
(** ntree **)  | 
15  | 
||
| 486 | 16  | 
goal Ntree.thy "ntree(A) = A * (UN n: nat. n -> ntree(A))";  | 
| 
529
 
f0d16216e394
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
17  | 
let open ntree; val rew = rewrite_rule con_defs in  | 
| 
 
f0d16216e394
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
18  | 
by (fast_tac (sum_cs addSIs (equalityI :: map rew intrs)  | 
| 
 
f0d16216e394
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
19  | 
addEs [rew elim]) 1)  | 
| 
 
f0d16216e394
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
20  | 
end;  | 
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
539 
diff
changeset
 | 
21  | 
qed "ntree_unfold";  | 
| 486 | 22  | 
|
23  | 
(*A nicer induction rule than the standard one*)  | 
|
24  | 
val major::prems = goal Ntree.thy  | 
|
| 1461 | 25  | 
"[| t: ntree(A); \  | 
| 486 | 26  | 
\ !!x n h. [| x: A; n: nat; h: n -> ntree(A); ALL i:n. P(h`i) \  | 
| 1461 | 27  | 
\ |] ==> P(Branch(x,h)) \  | 
| 486 | 28  | 
\ |] ==> P(t)";  | 
| 515 | 29  | 
by (rtac (major RS ntree.induct) 1);  | 
30  | 
by (etac UN_E 1);  | 
|
| 486 | 31  | 
by (REPEAT_SOME (ares_tac prems));  | 
32  | 
by (fast_tac (ZF_cs addEs [fun_weaken_type]) 1);  | 
|
33  | 
by (fast_tac (ZF_cs addDs [apply_type]) 1);  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
539 
diff
changeset
 | 
34  | 
qed "ntree_induct";  | 
| 486 | 35  | 
|
36  | 
(*Induction on ntree(A) to prove an equation*)  | 
|
37  | 
val major::prems = goal Ntree.thy  | 
|
| 1461 | 38  | 
"[| t: ntree(A); f: ntree(A)->B; g: ntree(A)->B; \  | 
| 486 | 39  | 
\ !!x n h. [| x: A; n: nat; h: n -> ntree(A); f O h = g O h |] ==> \  | 
| 1461 | 40  | 
\ f ` Branch(x,h) = g ` Branch(x,h) \  | 
| 486 | 41  | 
\ |] ==> f`t=g`t";  | 
| 539 | 42  | 
by (rtac (major RS ntree_induct) 1);  | 
| 486 | 43  | 
by (REPEAT_SOME (ares_tac prems));  | 
44  | 
by (cut_facts_tac prems 1);  | 
|
| 1461 | 45  | 
by (rtac fun_extension 1);  | 
| 486 | 46  | 
by (REPEAT_SOME (ares_tac [comp_fun]));  | 
47  | 
by (asm_simp_tac (ZF_ss addsimps [comp_fun_apply]) 1);  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
539 
diff
changeset
 | 
48  | 
qed "ntree_induct_eqn";  | 
| 486 | 49  | 
|
50  | 
(** Lemmas to justify using "Ntree" in other recursive type definitions **)  | 
|
51  | 
||
| 515 | 52  | 
goalw Ntree.thy ntree.defs "!!A B. A<=B ==> ntree(A) <= ntree(B)";  | 
| 486 | 53  | 
by (rtac lfp_mono 1);  | 
| 515 | 54  | 
by (REPEAT (rtac ntree.bnd_mono 1));  | 
| 486 | 55  | 
by (REPEAT (ares_tac (univ_mono::basic_monos) 1));  | 
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
539 
diff
changeset
 | 
56  | 
qed "ntree_mono";  | 
| 486 | 57  | 
|
58  | 
(*Easily provable by induction also*)  | 
|
| 515 | 59  | 
goalw Ntree.thy (ntree.defs@ntree.con_defs) "ntree(univ(A)) <= univ(A)";  | 
| 486 | 60  | 
by (rtac lfp_lowerbound 1);  | 
61  | 
by (rtac (A_subset_univ RS univ_mono) 2);  | 
|
62  | 
by (safe_tac ZF_cs);  | 
|
63  | 
by (REPEAT (ares_tac [Pair_in_univ, nat_fun_univ RS subsetD] 1));  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
539 
diff
changeset
 | 
64  | 
qed "ntree_univ";  | 
| 486 | 65  | 
|
| 539 | 66  | 
val ntree_subset_univ = [ntree_mono, ntree_univ] MRS subset_trans |> standard;  | 
67  | 
||
68  | 
||
69  | 
(** maptree **)  | 
|
70  | 
||
71  | 
goal Ntree.thy "maptree(A) = A * (maptree(A) -||> maptree(A))";  | 
|
72  | 
let open maptree; val rew = rewrite_rule con_defs in  | 
|
73  | 
by (fast_tac (sum_cs addSIs (equalityI :: map rew intrs)  | 
|
74  | 
addEs [rew elim]) 1)  | 
|
75  | 
end;  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
539 
diff
changeset
 | 
76  | 
qed "maptree_unfold";  | 
| 486 | 77  | 
|
| 539 | 78  | 
(*A nicer induction rule than the standard one*)  | 
79  | 
val major::prems = goal Ntree.thy  | 
|
| 1461 | 80  | 
"[| t: maptree(A); \  | 
81  | 
\ !!x n h. [| x: A; h: maptree(A) -||> maptree(A); \  | 
|
82  | 
\ ALL y: field(h). P(y) \  | 
|
83  | 
\ |] ==> P(Sons(x,h)) \  | 
|
| 539 | 84  | 
\ |] ==> P(t)";  | 
85  | 
by (rtac (major RS maptree.induct) 1);  | 
|
86  | 
by (REPEAT_SOME (ares_tac prems));  | 
|
87  | 
by (eresolve_tac [Collect_subset RS FiniteFun_mono1 RS subsetD] 1);  | 
|
88  | 
by (dresolve_tac [FiniteFun.dom_subset RS subsetD] 1);  | 
|
89  | 
by (dresolve_tac [Fin.dom_subset RS subsetD] 1);  | 
|
90  | 
by (fast_tac ZF_cs 1);  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
539 
diff
changeset
 | 
91  | 
qed "maptree_induct";  | 
| 539 | 92  | 
|
93  | 
||
94  | 
(** maptree2 **)  | 
|
95  | 
||
96  | 
goal Ntree.thy "maptree2(A,B) = A * (B -||> maptree2(A,B))";  | 
|
97  | 
let open maptree2; val rew = rewrite_rule con_defs in  | 
|
98  | 
by (fast_tac (sum_cs addSIs (equalityI :: map rew intrs)  | 
|
99  | 
addEs [rew elim]) 1)  | 
|
100  | 
end;  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
539 
diff
changeset
 | 
101  | 
qed "maptree2_unfold";  | 
| 539 | 102  | 
|
103  | 
(*A nicer induction rule than the standard one*)  | 
|
104  | 
val major::prems = goal Ntree.thy  | 
|
| 1461 | 105  | 
"[| t: maptree2(A,B); \  | 
| 539 | 106  | 
\ !!x n h. [| x: A; h: B -||> maptree2(A,B); ALL y:range(h). P(y) \  | 
| 1461 | 107  | 
\ |] ==> P(Sons2(x,h)) \  | 
| 539 | 108  | 
\ |] ==> P(t)";  | 
109  | 
by (rtac (major RS maptree2.induct) 1);  | 
|
110  | 
by (REPEAT_SOME (ares_tac prems));  | 
|
111  | 
by (eresolve_tac [[subset_refl, Collect_subset] MRS  | 
|
| 1461 | 112  | 
FiniteFun_mono RS subsetD] 1);  | 
| 539 | 113  | 
by (dresolve_tac [FiniteFun.dom_subset RS subsetD] 1);  | 
114  | 
by (dresolve_tac [Fin.dom_subset RS subsetD] 1);  | 
|
115  | 
by (fast_tac ZF_cs 1);  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
539 
diff
changeset
 | 
116  | 
qed "maptree2_induct";  | 
| 539 | 117  |