| 12201 |      1 | (*  Title:      ZF/Induct/Term.thy
 | 
|  |      2 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
|  |      3 |     Copyright   1994  University of Cambridge
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | header {* Terms over an alphabet *}
 | 
|  |      7 | 
 | 
| 16417 |      8 | theory Term imports Main begin
 | 
| 12201 |      9 | 
 | 
|  |     10 | text {*
 | 
|  |     11 |   Illustrates the list functor (essentially the same type as in @{text
 | 
|  |     12 |   Trees_Forest}).
 | 
|  |     13 | *}
 | 
|  |     14 | 
 | 
|  |     15 | consts
 | 
|  |     16 |   "term" :: "i => i"
 | 
|  |     17 | 
 | 
|  |     18 | datatype "term(A)" = Apply ("a \<in> A", "l \<in> list(term(A))")
 | 
|  |     19 |   monos list_mono
 | 
|  |     20 |   type_elims list_univ [THEN subsetD, elim_format]
 | 
|  |     21 | 
 | 
|  |     22 | declare Apply [TC]
 | 
|  |     23 | 
 | 
| 24893 |     24 | definition
 | 
|  |     25 |   term_rec :: "[i, [i, i, i] => i] => i"  where
 | 
| 12201 |     26 |   "term_rec(t,d) ==
 | 
|  |     27 |     Vrec(t, \<lambda>t g. term_case(\<lambda>x zs. d(x, zs, map(\<lambda>z. g`z, zs)), t))"
 | 
|  |     28 | 
 | 
| 24893 |     29 | definition
 | 
|  |     30 |   term_map :: "[i => i, i] => i"  where
 | 
| 12201 |     31 |   "term_map(f,t) == term_rec(t, \<lambda>x zs rs. Apply(f(x), rs))"
 | 
|  |     32 | 
 | 
| 24893 |     33 | definition
 | 
|  |     34 |   term_size :: "i => i"  where
 | 
| 12201 |     35 |   "term_size(t) == term_rec(t, \<lambda>x zs rs. succ(list_add(rs)))"
 | 
|  |     36 | 
 | 
| 24893 |     37 | definition
 | 
|  |     38 |   reflect :: "i => i"  where
 | 
| 12201 |     39 |   "reflect(t) == term_rec(t, \<lambda>x zs rs. Apply(x, rev(rs)))"
 | 
|  |     40 | 
 | 
| 24893 |     41 | definition
 | 
|  |     42 |   preorder :: "i => i"  where
 | 
| 12201 |     43 |   "preorder(t) == term_rec(t, \<lambda>x zs rs. Cons(x, flat(rs)))"
 | 
|  |     44 | 
 | 
| 24893 |     45 | definition
 | 
|  |     46 |   postorder :: "i => i"  where
 | 
| 12201 |     47 |   "postorder(t) == term_rec(t, \<lambda>x zs rs. flat(rs) @ [x])"
 | 
|  |     48 | 
 | 
|  |     49 | lemma term_unfold: "term(A) = A * list(term(A))"
 | 
|  |     50 |   by (fast intro!: term.intros [unfolded term.con_defs]
 | 
|  |     51 |     elim: term.cases [unfolded term.con_defs])
 | 
|  |     52 | 
 | 
|  |     53 | lemma term_induct2:
 | 
|  |     54 |   "[| t \<in> term(A);
 | 
|  |     55 |       !!x.      [| x \<in> A |] ==> P(Apply(x,Nil));
 | 
|  |     56 |       !!x z zs. [| x \<in> A;  z \<in> term(A);  zs: list(term(A));  P(Apply(x,zs))
 | 
|  |     57 |                 |] ==> P(Apply(x, Cons(z,zs)))
 | 
|  |     58 |      |] ==> P(t)"
 | 
|  |     59 |   -- {* Induction on @{term "term(A)"} followed by induction on @{term list}. *}
 | 
|  |     60 |   apply (induct_tac t)
 | 
|  |     61 |   apply (erule list.induct)
 | 
|  |     62 |    apply (auto dest: list_CollectD)
 | 
|  |     63 |   done
 | 
|  |     64 | 
 | 
| 18415 |     65 | lemma term_induct_eqn [consumes 1, case_names Apply]:
 | 
| 12201 |     66 |   "[| t \<in> term(A);
 | 
|  |     67 |       !!x zs. [| x \<in> A;  zs: list(term(A));  map(f,zs) = map(g,zs) |] ==>
 | 
|  |     68 |               f(Apply(x,zs)) = g(Apply(x,zs))
 | 
|  |     69 |    |] ==> f(t) = g(t)"
 | 
|  |     70 |   -- {* Induction on @{term "term(A)"} to prove an equation. *}
 | 
|  |     71 |   apply (induct_tac t)
 | 
|  |     72 |   apply (auto dest: map_list_Collect list_CollectD)
 | 
|  |     73 |   done
 | 
|  |     74 | 
 | 
|  |     75 | text {*
 | 
|  |     76 |   \medskip Lemmas to justify using @{term "term"} in other recursive
 | 
|  |     77 |   type definitions.
 | 
|  |     78 | *}
 | 
|  |     79 | 
 | 
|  |     80 | lemma term_mono: "A \<subseteq> B ==> term(A) \<subseteq> term(B)"
 | 
|  |     81 |   apply (unfold term.defs)
 | 
|  |     82 |   apply (rule lfp_mono)
 | 
|  |     83 |     apply (rule term.bnd_mono)+
 | 
|  |     84 |   apply (rule univ_mono basic_monos| assumption)+
 | 
|  |     85 |   done
 | 
|  |     86 | 
 | 
|  |     87 | lemma term_univ: "term(univ(A)) \<subseteq> univ(A)"
 | 
|  |     88 |   -- {* Easily provable by induction also *}
 | 
|  |     89 |   apply (unfold term.defs term.con_defs)
 | 
|  |     90 |   apply (rule lfp_lowerbound)
 | 
|  |     91 |    apply (rule_tac [2] A_subset_univ [THEN univ_mono])
 | 
|  |     92 |   apply safe
 | 
|  |     93 |   apply (assumption | rule Pair_in_univ list_univ [THEN subsetD])+
 | 
|  |     94 |   done
 | 
|  |     95 | 
 | 
|  |     96 | lemma term_subset_univ: "A \<subseteq> univ(B) ==> term(A) \<subseteq> univ(B)"
 | 
|  |     97 |   apply (rule subset_trans)
 | 
|  |     98 |    apply (erule term_mono)
 | 
|  |     99 |   apply (rule term_univ)
 | 
|  |    100 |   done
 | 
|  |    101 | 
 | 
|  |    102 | lemma term_into_univ: "[| t \<in> term(A);  A \<subseteq> univ(B) |] ==> t \<in> univ(B)"
 | 
|  |    103 |   by (rule term_subset_univ [THEN subsetD])
 | 
|  |    104 | 
 | 
|  |    105 | text {*
 | 
|  |    106 |   \medskip @{text term_rec} -- by @{text Vset} recursion.
 | 
|  |    107 | *}
 | 
|  |    108 | 
 | 
|  |    109 | lemma map_lemma: "[| l \<in> list(A);  Ord(i);  rank(l)<i |]
 | 
|  |    110 |     ==> map(\<lambda>z. (\<lambda>x \<in> Vset(i).h(x)) ` z, l) = map(h,l)"
 | 
|  |    111 |   -- {* @{term map} works correctly on the underlying list of terms. *}
 | 
|  |    112 |   apply (induct set: list)
 | 
|  |    113 |    apply simp
 | 
|  |    114 |   apply (subgoal_tac "rank (a) <i & rank (l) < i")
 | 
|  |    115 |    apply (simp add: rank_of_Ord)
 | 
|  |    116 |   apply (simp add: list.con_defs)
 | 
|  |    117 |   apply (blast dest: rank_rls [THEN lt_trans])
 | 
|  |    118 |   done
 | 
|  |    119 | 
 | 
|  |    120 | lemma term_rec [simp]: "ts \<in> list(A) ==>
 | 
|  |    121 |   term_rec(Apply(a,ts), d) = d(a, ts, map (\<lambda>z. term_rec(z,d), ts))"
 | 
|  |    122 |   -- {* Typing premise is necessary to invoke @{text map_lemma}. *}
 | 
|  |    123 |   apply (rule term_rec_def [THEN def_Vrec, THEN trans])
 | 
|  |    124 |   apply (unfold term.con_defs)
 | 
|  |    125 |   apply (simp add: rank_pair2 map_lemma)
 | 
|  |    126 |   done
 | 
|  |    127 | 
 | 
|  |    128 | lemma term_rec_type:
 | 
| 18415 |    129 |   assumes t: "t \<in> term(A)"
 | 
|  |    130 |     and a: "!!x zs r. [| x \<in> A;  zs: list(term(A));
 | 
| 12201 |    131 |                    r \<in> list(\<Union>t \<in> term(A). C(t)) |]
 | 
|  |    132 |                 ==> d(x, zs, r): C(Apply(x,zs))"
 | 
| 18415 |    133 |   shows "term_rec(t,d) \<in> C(t)"
 | 
|  |    134 |   -- {* Slightly odd typing condition on @{text r} in the second premise! *}
 | 
|  |    135 |   using t
 | 
|  |    136 |   apply induct
 | 
|  |    137 |   apply (frule list_CollectD)
 | 
|  |    138 |   apply (subst term_rec)
 | 
|  |    139 |    apply (assumption | rule a)+
 | 
|  |    140 |   apply (erule list.induct)
 | 
| 46900 |    141 |    apply auto
 | 
| 18415 |    142 |   done
 | 
| 12201 |    143 | 
 | 
|  |    144 | lemma def_term_rec:
 | 
|  |    145 |   "[| !!t. j(t)==term_rec(t,d);  ts: list(A) |] ==>
 | 
|  |    146 |     j(Apply(a,ts)) = d(a, ts, map(\<lambda>Z. j(Z), ts))"
 | 
|  |    147 |   apply (simp only:)
 | 
|  |    148 |   apply (erule term_rec)
 | 
|  |    149 |   done
 | 
|  |    150 | 
 | 
|  |    151 | lemma term_rec_simple_type [TC]:
 | 
|  |    152 |   "[| t \<in> term(A);
 | 
|  |    153 |       !!x zs r. [| x \<in> A;  zs: list(term(A));  r \<in> list(C) |]
 | 
|  |    154 |                 ==> d(x, zs, r): C
 | 
|  |    155 |    |] ==> term_rec(t,d) \<in> C"
 | 
|  |    156 |   apply (erule term_rec_type)
 | 
|  |    157 |   apply (drule subset_refl [THEN UN_least, THEN list_mono, THEN subsetD])
 | 
|  |    158 |   apply simp
 | 
|  |    159 |   done
 | 
|  |    160 | 
 | 
|  |    161 | 
 | 
|  |    162 | text {*
 | 
|  |    163 |   \medskip @{term term_map}.
 | 
|  |    164 | *}
 | 
|  |    165 | 
 | 
|  |    166 | lemma term_map [simp]:
 | 
| 12610 |    167 |   "ts \<in> list(A) ==>
 | 
|  |    168 |     term_map(f, Apply(a, ts)) = Apply(f(a), map(term_map(f), ts))"
 | 
| 12201 |    169 |   by (rule term_map_def [THEN def_term_rec])
 | 
|  |    170 | 
 | 
|  |    171 | lemma term_map_type [TC]:
 | 
|  |    172 |     "[| t \<in> term(A);  !!x. x \<in> A ==> f(x): B |] ==> term_map(f,t) \<in> term(B)"
 | 
|  |    173 |   apply (unfold term_map_def)
 | 
|  |    174 |   apply (erule term_rec_simple_type)
 | 
|  |    175 |   apply fast
 | 
|  |    176 |   done
 | 
|  |    177 | 
 | 
|  |    178 | lemma term_map_type2 [TC]:
 | 
|  |    179 |     "t \<in> term(A) ==> term_map(f,t) \<in> term({f(u). u \<in> A})"
 | 
|  |    180 |   apply (erule term_map_type)
 | 
|  |    181 |   apply (erule RepFunI)
 | 
|  |    182 |   done
 | 
|  |    183 | 
 | 
|  |    184 | text {*
 | 
|  |    185 |   \medskip @{term term_size}.
 | 
|  |    186 | *}
 | 
|  |    187 | 
 | 
|  |    188 | lemma term_size [simp]:
 | 
| 12216 |    189 |     "ts \<in> list(A) ==> term_size(Apply(a, ts)) = succ(list_add(map(term_size, ts)))"
 | 
| 12201 |    190 |   by (rule term_size_def [THEN def_term_rec])
 | 
|  |    191 | 
 | 
|  |    192 | lemma term_size_type [TC]: "t \<in> term(A) ==> term_size(t) \<in> nat"
 | 
|  |    193 |   by (auto simp add: term_size_def)
 | 
|  |    194 | 
 | 
|  |    195 | 
 | 
|  |    196 | text {*
 | 
|  |    197 |   \medskip @{text reflect}.
 | 
|  |    198 | *}
 | 
|  |    199 | 
 | 
|  |    200 | lemma reflect [simp]:
 | 
| 12216 |    201 |     "ts \<in> list(A) ==> reflect(Apply(a, ts)) = Apply(a, rev(map(reflect, ts)))"
 | 
| 12201 |    202 |   by (rule reflect_def [THEN def_term_rec])
 | 
|  |    203 | 
 | 
|  |    204 | lemma reflect_type [TC]: "t \<in> term(A) ==> reflect(t) \<in> term(A)"
 | 
|  |    205 |   by (auto simp add: reflect_def)
 | 
|  |    206 | 
 | 
|  |    207 | 
 | 
|  |    208 | text {*
 | 
|  |    209 |   \medskip @{text preorder}.
 | 
|  |    210 | *}
 | 
|  |    211 | 
 | 
|  |    212 | lemma preorder [simp]:
 | 
| 12216 |    213 |     "ts \<in> list(A) ==> preorder(Apply(a, ts)) = Cons(a, flat(map(preorder, ts)))"
 | 
| 12201 |    214 |   by (rule preorder_def [THEN def_term_rec])
 | 
|  |    215 | 
 | 
|  |    216 | lemma preorder_type [TC]: "t \<in> term(A) ==> preorder(t) \<in> list(A)"
 | 
|  |    217 |   by (simp add: preorder_def)
 | 
|  |    218 | 
 | 
|  |    219 | 
 | 
|  |    220 | text {*
 | 
|  |    221 |   \medskip @{text postorder}.
 | 
|  |    222 | *}
 | 
|  |    223 | 
 | 
|  |    224 | lemma postorder [simp]:
 | 
| 12216 |    225 |     "ts \<in> list(A) ==> postorder(Apply(a, ts)) = flat(map(postorder, ts)) @ [a]"
 | 
| 12201 |    226 |   by (rule postorder_def [THEN def_term_rec])
 | 
|  |    227 | 
 | 
|  |    228 | lemma postorder_type [TC]: "t \<in> term(A) ==> postorder(t) \<in> list(A)"
 | 
|  |    229 |   by (simp add: postorder_def)
 | 
|  |    230 | 
 | 
|  |    231 | 
 | 
|  |    232 | text {*
 | 
|  |    233 |   \medskip Theorems about @{text term_map}.
 | 
|  |    234 | *}
 | 
|  |    235 | 
 | 
| 26059 |    236 | declare map_compose [simp]
 | 
| 12201 |    237 | 
 | 
|  |    238 | lemma term_map_ident: "t \<in> term(A) ==> term_map(\<lambda>u. u, t) = t"
 | 
| 18415 |    239 |   by (induct rule: term_induct_eqn) simp
 | 
| 12201 |    240 | 
 | 
|  |    241 | lemma term_map_compose:
 | 
|  |    242 |     "t \<in> term(A) ==> term_map(f, term_map(g,t)) = term_map(\<lambda>u. f(g(u)), t)"
 | 
| 18415 |    243 |   by (induct rule: term_induct_eqn) simp
 | 
| 12201 |    244 | 
 | 
|  |    245 | lemma term_map_reflect:
 | 
|  |    246 |     "t \<in> term(A) ==> term_map(f, reflect(t)) = reflect(term_map(f,t))"
 | 
| 18415 |    247 |   by (induct rule: term_induct_eqn) (simp add: rev_map_distrib [symmetric])
 | 
| 12201 |    248 | 
 | 
|  |    249 | 
 | 
|  |    250 | text {*
 | 
|  |    251 |   \medskip Theorems about @{text term_size}.
 | 
|  |    252 | *}
 | 
|  |    253 | 
 | 
|  |    254 | lemma term_size_term_map: "t \<in> term(A) ==> term_size(term_map(f,t)) = term_size(t)"
 | 
| 18415 |    255 |   by (induct rule: term_induct_eqn) simp
 | 
| 12201 |    256 | 
 | 
|  |    257 | lemma term_size_reflect: "t \<in> term(A) ==> term_size(reflect(t)) = term_size(t)"
 | 
| 18415 |    258 |   by (induct rule: term_induct_eqn) (simp add: rev_map_distrib [symmetric] list_add_rev)
 | 
| 12201 |    259 | 
 | 
|  |    260 | lemma term_size_length: "t \<in> term(A) ==> term_size(t) = length(preorder(t))"
 | 
| 18415 |    261 |   by (induct rule: term_induct_eqn) (simp add: length_flat)
 | 
| 12201 |    262 | 
 | 
|  |    263 | 
 | 
|  |    264 | text {*
 | 
|  |    265 |   \medskip Theorems about @{text reflect}.
 | 
|  |    266 | *}
 | 
|  |    267 | 
 | 
|  |    268 | lemma reflect_reflect_ident: "t \<in> term(A) ==> reflect(reflect(t)) = t"
 | 
| 18415 |    269 |   by (induct rule: term_induct_eqn) (simp add: rev_map_distrib)
 | 
| 12201 |    270 | 
 | 
|  |    271 | 
 | 
|  |    272 | text {*
 | 
|  |    273 |   \medskip Theorems about preorder.
 | 
|  |    274 | *}
 | 
|  |    275 | 
 | 
| 12610 |    276 | lemma preorder_term_map:
 | 
|  |    277 |     "t \<in> term(A) ==> preorder(term_map(f,t)) = map(f, preorder(t))"
 | 
| 18415 |    278 |   by (induct rule: term_induct_eqn) (simp add: map_flat)
 | 
| 12201 |    279 | 
 | 
|  |    280 | lemma preorder_reflect_eq_rev_postorder:
 | 
|  |    281 |     "t \<in> term(A) ==> preorder(reflect(t)) = rev(postorder(t))"
 | 
| 18415 |    282 |   by (induct rule: term_induct_eqn)
 | 
|  |    283 |     (simp add: rev_app_distrib rev_flat rev_map_distrib [symmetric])
 | 
| 12201 |    284 | 
 | 
|  |    285 | end
 |