| 77573 |      1 | (* Author: Tobias Nipkow *)
 | 
|  |      2 | 
 | 
|  |      3 | section \<open>Tree Rotations\<close>
 | 
|  |      4 | 
 | 
|  |      5 | theory Tree_Rotations
 | 
|  |      6 | imports "HOL-Library.Tree"
 | 
|  |      7 | begin
 | 
|  |      8 | 
 | 
|  |      9 | text \<open>How to transform a tree into a list and into any other tree (with the same @{const inorder})
 | 
|  |     10 | by rotations.\<close>
 | 
|  |     11 | 
 | 
|  |     12 | fun is_list :: "'a tree \<Rightarrow> bool" where
 | 
|  |     13 | "is_list (Node l _ r) = (l = Leaf \<and> is_list r)" |
 | 
|  |     14 | "is_list Leaf = True"
 | 
|  |     15 | 
 | 
|  |     16 | text \<open>Termination proof via measure function. NB @{term "size t - rlen t"} works for
 | 
|  |     17 | the actual rotation equation but not for the second equation.\<close>
 | 
|  |     18 | 
 | 
|  |     19 | fun rlen :: "'a tree \<Rightarrow> nat" where
 | 
|  |     20 | "rlen Leaf = 0" |
 | 
|  |     21 | "rlen (Node l x r) = rlen r + 1"
 | 
|  |     22 | 
 | 
|  |     23 | lemma rlen_le_size: "rlen t \<le> size t"
 | 
|  |     24 | by(induction t) auto
 | 
|  |     25 | 
 | 
| 77575 |     26 | 
 | 
|  |     27 | subsection \<open>Without positions\<close>
 | 
|  |     28 | 
 | 
|  |     29 | function (sequential) list_of :: "'a tree \<Rightarrow> 'a tree" where
 | 
|  |     30 | "list_of (Node (Node A a B) b C) = list_of (Node A a (Node B b C))" |
 | 
|  |     31 | "list_of (Node Leaf a A) = Node Leaf a (list_of A)" |
 | 
|  |     32 | "list_of Leaf = Leaf"
 | 
| 77573 |     33 | by pat_completeness auto
 | 
|  |     34 | 
 | 
|  |     35 | termination
 | 
|  |     36 | proof
 | 
|  |     37 |   let ?R = "measure(\<lambda>t. 2*size t - rlen t)"
 | 
|  |     38 |   show "wf ?R" by (auto simp add: mlex_prod_def)
 | 
|  |     39 | 
 | 
|  |     40 |   fix A a B b C
 | 
|  |     41 |   show "(Node A a (Node B b C), Node (Node A a B) b C) \<in> ?R"
 | 
|  |     42 |     using rlen_le_size[of C] by(simp)
 | 
|  |     43 | 
 | 
|  |     44 |   fix a A show "(A, Node Leaf a A) \<in> ?R" using rlen_le_size[of A] by(simp)
 | 
|  |     45 | qed
 | 
|  |     46 | 
 | 
| 77575 |     47 | lemma is_list_rot: "is_list(list_of t)"
 | 
|  |     48 | by (induction t rule: list_of.induct) auto
 | 
| 77573 |     49 | 
 | 
| 77575 |     50 | lemma inorder_rot: "inorder(list_of t) = inorder t"
 | 
|  |     51 | by (induction t rule: list_of.induct) auto
 | 
| 77573 |     52 | 
 | 
|  |     53 | 
 | 
| 77575 |     54 | subsection \<open>With positions\<close>
 | 
| 77573 |     55 | 
 | 
|  |     56 | datatype dir = L | R
 | 
|  |     57 | 
 | 
|  |     58 | type_synonym "pos" = "dir list"
 | 
|  |     59 | 
 | 
|  |     60 | function (sequential) rotR_poss :: "'a tree \<Rightarrow> pos list" where
 | 
|  |     61 | "rotR_poss (Node (Node A a B) b C) = [] # rotR_poss (Node A a (Node B b C))" |
 | 
|  |     62 | "rotR_poss (Node Leaf a A) = map (Cons R) (rotR_poss A)" |
 | 
|  |     63 | "rotR_poss Leaf = []"
 | 
|  |     64 | by pat_completeness auto
 | 
|  |     65 | 
 | 
|  |     66 | termination
 | 
|  |     67 | proof
 | 
|  |     68 |   let ?R = "measure(\<lambda>t. 2*size t - rlen t)"
 | 
|  |     69 |   show "wf ?R" by (auto simp add: mlex_prod_def)
 | 
|  |     70 | 
 | 
|  |     71 |   fix A a B b C
 | 
|  |     72 |   show "(Node A a (Node B b C), Node (Node A a B) b C) \<in> ?R"
 | 
|  |     73 |     using rlen_le_size[of C] by(simp)
 | 
|  |     74 | 
 | 
|  |     75 |   fix a A show "(A, Node Leaf a A) \<in> ?R" using rlen_le_size[of A] by(simp)
 | 
|  |     76 | qed
 | 
|  |     77 | 
 | 
|  |     78 | fun rotR :: "'a tree \<Rightarrow> 'a tree" where
 | 
|  |     79 | "rotR (Node (Node A a B) b C) = Node A a (Node B b C)"
 | 
|  |     80 | 
 | 
|  |     81 | fun rotL :: "'a tree \<Rightarrow> 'a tree" where
 | 
|  |     82 | "rotL (Node A a (Node B b C)) = Node (Node A a B) b C"
 | 
|  |     83 | 
 | 
|  |     84 | fun apply_at :: "('a tree \<Rightarrow> 'a tree) \<Rightarrow> pos \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
 | 
|  |     85 |   "apply_at f [] t = f t"
 | 
|  |     86 | | "apply_at f (L # ds) (Node l a r) = Node (apply_at f ds l) a r"
 | 
|  |     87 | | "apply_at f (R # ds) (Node l a r) = Node l a (apply_at f ds r)"
 | 
|  |     88 | 
 | 
|  |     89 | fun apply_ats :: "('a tree \<Rightarrow> 'a tree) \<Rightarrow> pos list \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
 | 
|  |     90 | "apply_ats _ [] t = t" |
 | 
|  |     91 | "apply_ats f (p#ps) t = apply_ats f ps (apply_at f p t)"
 | 
|  |     92 | 
 | 
|  |     93 | lemma apply_ats_append:
 | 
|  |     94 |   "apply_ats f (ps\<^sub>1 @ ps\<^sub>2) t = apply_ats f ps\<^sub>2 (apply_ats f ps\<^sub>1 t)"
 | 
|  |     95 | by (induction ps\<^sub>1 arbitrary: t) auto
 | 
|  |     96 | 
 | 
|  |     97 | abbreviation "rotRs \<equiv> apply_ats rotR"
 | 
|  |     98 | abbreviation "rotLs \<equiv> apply_ats rotL"
 | 
|  |     99 | 
 | 
|  |    100 | lemma apply_ats_map_R: "apply_ats f (map ((#) R) ps) \<langle>l, a, r\<rangle> = Node l a (apply_ats f ps r)"
 | 
|  |    101 | by(induction ps arbitrary: r) auto
 | 
|  |    102 | 
 | 
|  |    103 | lemma inorder_rotRs_poss: "inorder (rotRs (rotR_poss t) t) = inorder t"
 | 
|  |    104 | apply(induction t rule: rotR_poss.induct)
 | 
|  |    105 | apply(auto simp: apply_ats_map_R)
 | 
|  |    106 | done
 | 
|  |    107 | 
 | 
|  |    108 | lemma is_list_rotRs: "is_list (rotRs (rotR_poss t) t)"
 | 
|  |    109 | apply(induction t rule: rotR_poss.induct)
 | 
|  |    110 | apply(auto simp: apply_ats_map_R)
 | 
|  |    111 | done
 | 
|  |    112 | 
 | 
|  |    113 | lemma "is_list (rotRs ps t) \<longrightarrow> length ps \<le> length(rotR_poss t)"
 | 
|  |    114 | quickcheck[expect=counterexample]
 | 
|  |    115 | oops
 | 
|  |    116 | 
 | 
|  |    117 | lemma length_rotRs_poss: "length (rotR_poss t) = size t - rlen t"
 | 
|  |    118 | proof(induction t rule: rotR_poss.induct)
 | 
|  |    119 |   case (1 A a B b C)
 | 
|  |    120 |   then show ?case using rlen_le_size[of C] by simp
 | 
|  |    121 | qed auto
 | 
|  |    122 | 
 | 
|  |    123 | lemma is_list_inorder_same:
 | 
|  |    124 |   "is_list t1 \<Longrightarrow> is_list t2 \<Longrightarrow> inorder t1 = inorder t2 \<Longrightarrow> t1 = t2"
 | 
|  |    125 | proof(induction t1 arbitrary: t2)
 | 
|  |    126 |   case Leaf
 | 
|  |    127 |   then show ?case by simp
 | 
|  |    128 | next
 | 
|  |    129 |   case Node
 | 
|  |    130 |   then show ?case by (cases t2) simp_all
 | 
|  |    131 | qed
 | 
|  |    132 | 
 | 
|  |    133 | lemma rot_id: "rotLs (rev (rotR_poss t)) (rotRs (rotR_poss t) t) = t"
 | 
|  |    134 | apply(induction t rule: rotR_poss.induct)
 | 
|  |    135 | apply(auto simp: apply_ats_map_R rev_map apply_ats_append)
 | 
|  |    136 | done
 | 
|  |    137 | 
 | 
|  |    138 | corollary tree_to_tree_rotations: assumes "inorder t1 = inorder t2"
 | 
|  |    139 | shows "rotLs (rev (rotR_poss t2)) (rotRs (rotR_poss t1) t1) = t2"
 | 
|  |    140 | proof -
 | 
|  |    141 |   have "rotRs (rotR_poss t1) t1 = rotRs (rotR_poss t2) t2" (is "?L = ?R")
 | 
|  |    142 |     by (simp add: assms inorder_rotRs_poss is_list_inorder_same is_list_rotRs)
 | 
|  |    143 |   hence "rotLs (rev (rotR_poss t2)) ?L = rotLs (rev (rotR_poss t2)) ?R"
 | 
|  |    144 |     by simp
 | 
|  |    145 |   also have "\<dots> = t2" by(rule rot_id)
 | 
|  |    146 |   finally show ?thesis .
 | 
|  |    147 | qed
 | 
|  |    148 | 
 | 
|  |    149 | lemma size_rlen_better_ub: "size t - rlen t \<le> size t - 1"
 | 
|  |    150 | by (cases t) auto
 | 
|  |    151 | 
 | 
|  |    152 | end
 |