| 59928 |      1 | (* Author: Tobias Nipkow *)
 | 
|  |      2 | 
 | 
| 60500 |      3 | section \<open>Multiset of Elements of Binary Tree\<close>
 | 
| 59928 |      4 | 
 | 
|  |      5 | theory Tree_Multiset
 | 
|  |      6 | imports Multiset Tree
 | 
|  |      7 | begin
 | 
|  |      8 | 
 | 
| 60500 |      9 | text\<open>Kept separate from theory @{theory Tree} to avoid importing all of
 | 
| 59928 |     10 | theory @{theory Multiset} into @{theory Tree}. Should be merged if
 | 
| 60500 |     11 | @{theory Multiset} ever becomes part of @{theory Main}.\<close>
 | 
| 59928 |     12 | 
 | 
|  |     13 | fun mset_tree :: "'a tree \<Rightarrow> 'a multiset" where
 | 
|  |     14 | "mset_tree Leaf = {#}" |
 | 
|  |     15 | "mset_tree (Node l a r) = {#a#} + mset_tree l + mset_tree r"
 | 
|  |     16 | 
 | 
| 60495 |     17 | lemma set_mset_tree[simp]: "set_mset (mset_tree t) = set_tree t"
 | 
| 59928 |     18 | by(induction t) auto
 | 
|  |     19 | 
 | 
|  |     20 | lemma size_mset_tree[simp]: "size(mset_tree t) = size t"
 | 
|  |     21 | by(induction t) auto
 | 
|  |     22 | 
 | 
|  |     23 | lemma mset_map_tree: "mset_tree (map_tree f t) = image_mset f (mset_tree t)"
 | 
|  |     24 | by (induction t) auto
 | 
|  |     25 | 
 | 
| 60505 |     26 | lemma mset_iff_set_tree: "x \<in># mset_tree t \<longleftrightarrow> x \<in> set_tree t"
 | 
|  |     27 | by(induction t arbitrary: x) auto
 | 
|  |     28 | 
 | 
| 60515 |     29 | lemma mset_preorder[simp]: "mset (preorder t) = mset_tree t"
 | 
| 59928 |     30 | by (induction t) (auto simp: ac_simps)
 | 
|  |     31 | 
 | 
| 60515 |     32 | lemma mset_inorder[simp]: "mset (inorder t) = mset_tree t"
 | 
| 59928 |     33 | by (induction t) (auto simp: ac_simps)
 | 
|  |     34 | 
 | 
|  |     35 | lemma map_mirror: "mset_tree (mirror t) = mset_tree t"
 | 
|  |     36 | by (induction t) (simp_all add: ac_simps)
 | 
|  |     37 | 
 | 
|  |     38 | end
 |