src/HOL/Library/Tree_Multiset.thy
author nipkow
Mon, 06 Apr 2015 15:23:50 +0200
changeset 59928 b9b7f913a19a
child 60495 d7ff0a1df90a
child 60500 903bb1495239
permissions -rw-r--r--
new theory Library/Tree_Multiset.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
59928
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
     2
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
     3
section {* Multiset of Elements of Binary Tree *}
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
     4
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
     5
theory Tree_Multiset
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
     6
imports Multiset Tree
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
     7
begin
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
     8
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
     9
text{* Kept separate from theory @{theory Tree} to avoid importing all of
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    10
theory @{theory Multiset} into @{theory Tree}. Should be merged if
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    11
@{theory Multiset} ever becomes part of @{theory Main}. *}
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    12
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    13
fun mset_tree :: "'a tree \<Rightarrow> 'a multiset" where
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    14
"mset_tree Leaf = {#}" |
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    15
"mset_tree (Node l a r) = {#a#} + mset_tree l + mset_tree r"
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    16
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    17
lemma set_of_mset_tree[simp]: "set_of (mset_tree t) = set_tree t"
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    18
by(induction t) auto
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    19
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    20
lemma size_mset_tree[simp]: "size(mset_tree t) = size t"
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    21
by(induction t) auto
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    22
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    23
lemma mset_map_tree: "mset_tree (map_tree f t) = image_mset f (mset_tree t)"
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    24
by (induction t) auto
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    25
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    26
lemma multiset_of_preorder[simp]: "multiset_of (preorder t) = mset_tree t"
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    27
by (induction t) (auto simp: ac_simps)
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    28
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    29
lemma multiset_of_inorder[simp]: "multiset_of (inorder t) = mset_tree t"
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    30
by (induction t) (auto simp: ac_simps)
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    31
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    32
lemma map_mirror: "mset_tree (mirror t) = mset_tree t"
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    33
by (induction t) (simp_all add: ac_simps)
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    34
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    35
lemma del_rightmost_mset_tree:
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    36
  "\<lbrakk> del_rightmost t = (t',x);  t \<noteq> \<langle>\<rangle> \<rbrakk> \<Longrightarrow> mset_tree t = {#x#} + mset_tree t'"
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    37
apply(induction t arbitrary: t' rule: del_rightmost.induct)
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    38
by (auto split: prod.splits) (auto simp: ac_simps)
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    39
b9b7f913a19a new theory Library/Tree_Multiset.thy
nipkow
parents:
diff changeset
    40
end