| author | wenzelm | 
| Sun, 17 Jan 2021 23:48:55 +0100 | |
| changeset 73145 | 661e9bc0411e | 
| parent 69593 | 3dda49e08b9d | 
| child 76213 | e44d86131648 | 
| permissions | -rw-r--r-- | 
| 12229 | 1  | 
(* Title: ZF/Induct/Ntree.thy  | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
3  | 
Copyright 1994 University of Cambridge  | 
|
4  | 
*)  | 
|
5  | 
||
| 60770 | 6  | 
section \<open>Datatype definition n-ary branching trees\<close>  | 
| 12229 | 7  | 
|
| 
65449
 
c82e63b11b8b
clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
 
wenzelm 
parents: 
61798 
diff
changeset
 | 
8  | 
theory Ntree imports ZF begin  | 
| 12229 | 9  | 
|
| 60770 | 10  | 
text \<open>  | 
| 12229 | 11  | 
Demonstrates a simple use of function space in a datatype  | 
| 61798 | 12  | 
definition. Based upon theory \<open>Term\<close>.  | 
| 60770 | 13  | 
\<close>  | 
| 12229 | 14  | 
|
15  | 
consts  | 
|
16  | 
ntree :: "i => i"  | 
|
17  | 
maptree :: "i => i"  | 
|
18  | 
maptree2 :: "[i, i] => i"  | 
|
19  | 
||
20  | 
datatype "ntree(A)" = Branch ("a \<in> A", "h \<in> (\<Union>n \<in> nat. n -> ntree(A))")
 | 
|
| 61798 | 21  | 
monos UN_mono [OF subset_refl Pi_mono] \<comment> \<open>MUST have this form\<close>  | 
| 12229 | 22  | 
type_intros nat_fun_univ [THEN subsetD]  | 
23  | 
type_elims UN_E  | 
|
24  | 
||
25  | 
datatype "maptree(A)" = Sons ("a \<in> A", "h \<in> maptree(A) -||> maptree(A)")
 | 
|
| 61798 | 26  | 
monos FiniteFun_mono1 \<comment> \<open>Use monotonicity in BOTH args\<close>  | 
| 12229 | 27  | 
type_intros FiniteFun_univ1 [THEN subsetD]  | 
28  | 
||
29  | 
datatype "maptree2(A, B)" = Sons2 ("a \<in> A", "h \<in> B -||> maptree2(A, B)")
 | 
|
30  | 
monos FiniteFun_mono [OF subset_refl]  | 
|
31  | 
type_intros FiniteFun_in_univ'  | 
|
32  | 
||
| 24893 | 33  | 
definition  | 
34  | 
ntree_rec :: "[[i, i, i] => i, i] => i" where  | 
|
| 12229 | 35  | 
"ntree_rec(b) ==  | 
| 12610 | 36  | 
Vrecursor(\<lambda>pr. ntree_case(\<lambda>x h. b(x, h, \<lambda>i \<in> domain(h). pr`(h`i))))"  | 
| 12229 | 37  | 
|
| 24893 | 38  | 
definition  | 
39  | 
ntree_copy :: "i => i" where  | 
|
| 12610 | 40  | 
"ntree_copy(z) == ntree_rec(\<lambda>x h r. Branch(x,r), z)"  | 
| 12229 | 41  | 
|
42  | 
||
| 60770 | 43  | 
text \<open>  | 
| 61798 | 44  | 
\medskip \<open>ntree\<close>  | 
| 60770 | 45  | 
\<close>  | 
| 12229 | 46  | 
|
47  | 
lemma ntree_unfold: "ntree(A) = A \<times> (\<Union>n \<in> nat. n -> ntree(A))"  | 
|
| 12788 | 48  | 
by (blast intro: ntree.intros [unfolded ntree.con_defs]  | 
| 12229 | 49  | 
elim: ntree.cases [unfolded ntree.con_defs])  | 
50  | 
||
| 18372 | 51  | 
lemma ntree_induct [consumes 1, case_names Branch, induct set: ntree]:  | 
52  | 
assumes t: "t \<in> ntree(A)"  | 
|
53  | 
and step: "!!x n h. [| x \<in> A; n \<in> nat; h \<in> n -> ntree(A); \<forall>i \<in> n. P(h`i)  | 
|
54  | 
|] ==> P(Branch(x,h))"  | 
|
55  | 
shows "P(t)"  | 
|
| 61798 | 56  | 
\<comment> \<open>A nicer induction rule than the standard one.\<close>  | 
| 18372 | 57  | 
using t  | 
58  | 
apply induct  | 
|
59  | 
apply (erule UN_E)  | 
|
60  | 
apply (assumption | rule step)+  | 
|
61  | 
apply (fast elim: fun_weaken_type)  | 
|
62  | 
apply (fast dest: apply_type)  | 
|
63  | 
done  | 
|
| 12229 | 64  | 
|
| 18372 | 65  | 
lemma ntree_induct_eqn [consumes 1]:  | 
66  | 
assumes t: "t \<in> ntree(A)"  | 
|
67  | 
and f: "f \<in> ntree(A)->B"  | 
|
68  | 
and g: "g \<in> ntree(A)->B"  | 
|
69  | 
and step: "!!x n h. [| x \<in> A; n \<in> nat; h \<in> n -> ntree(A); f O h = g O h |] ==>  | 
|
70  | 
f ` Branch(x,h) = g ` Branch(x,h)"  | 
|
71  | 
shows "f`t=g`t"  | 
|
| 69593 | 72  | 
\<comment> \<open>Induction on \<^term>\<open>ntree(A)\<close> to prove an equation\<close>  | 
| 18372 | 73  | 
using t  | 
74  | 
apply induct  | 
|
75  | 
apply (assumption | rule step)+  | 
|
76  | 
apply (insert f g)  | 
|
77  | 
apply (rule fun_extension)  | 
|
78  | 
apply (assumption | rule comp_fun)+  | 
|
79  | 
apply (simp add: comp_fun_apply)  | 
|
| 12229 | 80  | 
done  | 
81  | 
||
| 60770 | 82  | 
text \<open>  | 
| 61798 | 83  | 
\medskip Lemmas to justify using \<open>Ntree\<close> in other recursive  | 
| 12229 | 84  | 
type definitions.  | 
| 60770 | 85  | 
\<close>  | 
| 12229 | 86  | 
|
87  | 
lemma ntree_mono: "A \<subseteq> B ==> ntree(A) \<subseteq> ntree(B)"  | 
|
88  | 
apply (unfold ntree.defs)  | 
|
89  | 
apply (rule lfp_mono)  | 
|
90  | 
apply (rule ntree.bnd_mono)+  | 
|
91  | 
apply (assumption | rule univ_mono basic_monos)+  | 
|
92  | 
done  | 
|
93  | 
||
94  | 
lemma ntree_univ: "ntree(univ(A)) \<subseteq> univ(A)"  | 
|
| 61798 | 95  | 
\<comment> \<open>Easily provable by induction also\<close>  | 
| 12229 | 96  | 
apply (unfold ntree.defs ntree.con_defs)  | 
97  | 
apply (rule lfp_lowerbound)  | 
|
98  | 
apply (rule_tac [2] A_subset_univ [THEN univ_mono])  | 
|
99  | 
apply (blast intro: Pair_in_univ nat_fun_univ [THEN subsetD])  | 
|
100  | 
done  | 
|
101  | 
||
102  | 
lemma ntree_subset_univ: "A \<subseteq> univ(B) ==> ntree(A) \<subseteq> univ(B)"  | 
|
103  | 
by (rule subset_trans [OF ntree_mono ntree_univ])  | 
|
104  | 
||
105  | 
||
| 60770 | 106  | 
text \<open>  | 
| 61798 | 107  | 
\medskip \<open>ntree\<close> recursion.  | 
| 60770 | 108  | 
\<close>  | 
| 12229 | 109  | 
|
| 
13175
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12788 
diff
changeset
 | 
110  | 
lemma ntree_rec_Branch:  | 
| 
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12788 
diff
changeset
 | 
111  | 
"function(h) ==>  | 
| 
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12788 
diff
changeset
 | 
112  | 
ntree_rec(b, Branch(x,h)) = b(x, h, \<lambda>i \<in> domain(h). ntree_rec(b, h`i))"  | 
| 12229 | 113  | 
apply (rule ntree_rec_def [THEN def_Vrecursor, THEN trans])  | 
114  | 
apply (simp add: ntree.con_defs rank_pair2 [THEN [2] lt_trans] rank_apply)  | 
|
115  | 
done  | 
|
116  | 
||
117  | 
lemma ntree_copy_Branch [simp]:  | 
|
| 
13175
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12788 
diff
changeset
 | 
118  | 
"function(h) ==>  | 
| 
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12788 
diff
changeset
 | 
119  | 
ntree_copy (Branch(x, h)) = Branch(x, \<lambda>i \<in> domain(h). ntree_copy (h`i))"  | 
| 
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12788 
diff
changeset
 | 
120  | 
by (simp add: ntree_copy_def ntree_rec_Branch)  | 
| 12229 | 121  | 
|
122  | 
lemma ntree_copy_is_ident: "z \<in> ntree(A) ==> ntree_copy(z) = z"  | 
|
| 18372 | 123  | 
by (induct z set: ntree)  | 
124  | 
(auto simp add: domain_of_fun Pi_Collect_iff fun_is_function)  | 
|
| 12229 | 125  | 
|
126  | 
||
| 60770 | 127  | 
text \<open>  | 
| 61798 | 128  | 
\medskip \<open>maptree\<close>  | 
| 60770 | 129  | 
\<close>  | 
| 12229 | 130  | 
|
131  | 
lemma maptree_unfold: "maptree(A) = A \<times> (maptree(A) -||> maptree(A))"  | 
|
132  | 
by (fast intro!: maptree.intros [unfolded maptree.con_defs]  | 
|
133  | 
elim: maptree.cases [unfolded maptree.con_defs])  | 
|
134  | 
||
| 18372 | 135  | 
lemma maptree_induct [consumes 1, induct set: maptree]:  | 
136  | 
assumes t: "t \<in> maptree(A)"  | 
|
137  | 
and step: "!!x n h. [| x \<in> A; h \<in> maptree(A) -||> maptree(A);  | 
|
| 12229 | 138  | 
\<forall>y \<in> field(h). P(y)  | 
| 18372 | 139  | 
|] ==> P(Sons(x,h))"  | 
140  | 
shows "P(t)"  | 
|
| 61798 | 141  | 
\<comment> \<open>A nicer induction rule than the standard one.\<close>  | 
| 18372 | 142  | 
using t  | 
143  | 
apply induct  | 
|
144  | 
apply (assumption | rule step)+  | 
|
145  | 
apply (erule Collect_subset [THEN FiniteFun_mono1, THEN subsetD])  | 
|
146  | 
apply (drule FiniteFun.dom_subset [THEN subsetD])  | 
|
147  | 
apply (drule Fin.dom_subset [THEN subsetD])  | 
|
148  | 
apply fast  | 
|
149  | 
done  | 
|
| 12229 | 150  | 
|
151  | 
||
| 60770 | 152  | 
text \<open>  | 
| 61798 | 153  | 
\medskip \<open>maptree2\<close>  | 
| 60770 | 154  | 
\<close>  | 
| 12229 | 155  | 
|
156  | 
lemma maptree2_unfold: "maptree2(A, B) = A \<times> (B -||> maptree2(A, B))"  | 
|
157  | 
by (fast intro!: maptree2.intros [unfolded maptree2.con_defs]  | 
|
158  | 
elim: maptree2.cases [unfolded maptree2.con_defs])  | 
|
159  | 
||
| 18372 | 160  | 
lemma maptree2_induct [consumes 1, induct set: maptree2]:  | 
161  | 
assumes t: "t \<in> maptree2(A, B)"  | 
|
162  | 
and step: "!!x n h. [| x \<in> A; h \<in> B -||> maptree2(A,B); \<forall>y \<in> range(h). P(y)  | 
|
163  | 
|] ==> P(Sons2(x,h))"  | 
|
164  | 
shows "P(t)"  | 
|
165  | 
using t  | 
|
166  | 
apply induct  | 
|
167  | 
apply (assumption | rule step)+  | 
|
168  | 
apply (erule FiniteFun_mono [OF subset_refl Collect_subset, THEN subsetD])  | 
|
169  | 
apply (drule FiniteFun.dom_subset [THEN subsetD])  | 
|
170  | 
apply (drule Fin.dom_subset [THEN subsetD])  | 
|
171  | 
apply fast  | 
|
172  | 
done  | 
|
| 12229 | 173  | 
|
174  | 
end  |