author | wenzelm |
Mon, 13 Apr 2020 22:08:14 +0200 | |
changeset 71751 | abf3e80bd815 |
parent 70762 | d4a23cc9aabc |
child 71796 | 641f4c8ffec8 |
permissions | -rw-r--r-- |
61224 | 1 |
theory Tree2 |
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
2 |
imports "HOL-Library.Tree" |
61224 | 3 |
begin |
4 |
||
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
5 |
text \<open>This theory provides the basic infrastructure for the type @{typ \<open>('a * 'b) tree\<close>} |
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
6 |
of augmented trees where @{typ 'a} is the key and @{typ 'b} some additional information.\<close> |
61224 | 7 |
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
8 |
text \<open>IMPORTANT: Inductions and cases analyses on augmented trees need to use the following |
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
9 |
two rules explicitly. They generate nodes of the form @{term "Node l (a,b) r"} |
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
10 |
rather than @{term "Node l a r"} for trees of type @{typ "'a tree"}.\<close> |
61224 | 11 |
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
12 |
lemmas tree2_induct = tree.induct[where 'a = "'a * 'b", split_format(complete)] |
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
13 |
|
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
14 |
lemmas tree2_cases = tree.exhaust[where 'a = "'a * 'b", split_format(complete)] |
67967 | 15 |
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
16 |
fun inorder :: "('a*'b)tree \<Rightarrow> 'a list" where |
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
17 |
"inorder Leaf = []" | |
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
18 |
"inorder (Node l (a,_) r) = inorder l @ a # inorder r" |
62650 | 19 |
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
20 |
fun set_tree :: "('a*'b) tree \<Rightarrow> 'a set" where |
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
21 |
"set_tree Leaf = {}" | |
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
22 |
"set_tree (Node l (a,_) r) = Set.insert a (set_tree l \<union> set_tree r)" |
70742 | 23 |
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
24 |
fun bst :: "('a::linorder*'b) tree \<Rightarrow> bool" where |
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
25 |
"bst Leaf = True" | |
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70745
diff
changeset
|
26 |
"bst (Node l (a, _) r) = (bst l \<and> bst r \<and> (\<forall>x \<in> set_tree l. x < a) \<and> (\<forall>x \<in> set_tree r. a < x))" |
62650 | 27 |
|
68411 | 28 |
lemma finite_set_tree[simp]: "finite(set_tree t)" |
29 |
by(induction t) auto |
|
30 |
||
70745 | 31 |
lemma eq_set_tree_empty[simp]: "set_tree t = {} \<longleftrightarrow> t = Leaf" |
70744 | 32 |
by (cases t) auto |
33 |
||
70762 | 34 |
lemma set_inorder[simp]: "set (inorder t) = set_tree t" |
35 |
by (induction t) auto |
|
36 |
||
62390 | 37 |
end |