src/HOL/Data_Structures/Tree_Rotations.thy
changeset 77573 237e5504bae7
child 77575 72a99b54e206
equal deleted inserted replaced
77552:080422b3d914 77573:237e5504bae7
       
     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 
       
    26 function (sequential) rot_to_list :: "'a tree \<Rightarrow> 'a tree" where
       
    27 "rot_to_list (Node (Node A a B) b C) = rot_to_list (Node A a (Node B b C))" |
       
    28 "rot_to_list (Node Leaf a A) = Node Leaf a (rot_to_list A)" |
       
    29 "rot_to_list Leaf = Leaf"
       
    30 by pat_completeness auto
       
    31 
       
    32 termination
       
    33 proof
       
    34   let ?R = "measure(\<lambda>t. 2*size t - rlen t)"
       
    35   show "wf ?R" by (auto simp add: mlex_prod_def)
       
    36 
       
    37   fix A a B b C
       
    38   show "(Node A a (Node B b C), Node (Node A a B) b C) \<in> ?R"
       
    39     using rlen_le_size[of C] by(simp)
       
    40 
       
    41   fix a A show "(A, Node Leaf a A) \<in> ?R" using rlen_le_size[of A] by(simp)
       
    42 qed
       
    43 
       
    44 lemma is_list_rot: "is_list(rot_to_list t)"
       
    45 by (induction t rule: rot_to_list.induct) auto
       
    46 
       
    47 lemma inorder_rot: "inorder(rot_to_list t) = inorder t"
       
    48 by (induction t rule: rot_to_list.induct) auto
       
    49 
       
    50 function (sequential) n_rot_to_list :: "'a tree \<Rightarrow> nat" where
       
    51 "n_rot_to_list (Node (Node A a B) b C) = n_rot_to_list (Node A a (Node B b C)) + 1" |
       
    52 "n_rot_to_list (Node Leaf a A) = n_rot_to_list A" |
       
    53 "n_rot_to_list Leaf = 0"
       
    54 by pat_completeness auto
       
    55 
       
    56 termination
       
    57 proof
       
    58   let ?R = "measure(\<lambda>t. 2*size t - rlen t)"
       
    59   show "wf ?R" by (auto simp add: mlex_prod_def)
       
    60 
       
    61   fix A a B b C
       
    62   show "(Node A a (Node B b C), Node (Node A a B) b C) \<in> ?R"
       
    63     using rlen_le_size[of C] by(simp)
       
    64 
       
    65   fix a A show "(A, Node Leaf a A) \<in> ?R" using rlen_le_size[of A] by(simp)
       
    66 qed
       
    67 
       
    68 text \<open>Closed formula for the exact number of rotations needed:\<close>
       
    69 
       
    70 lemma n_rot_to_list: "n_rot_to_list t = size t - rlen t"
       
    71 proof(induction t rule: n_rot_to_list.induct)
       
    72   case (1 A a B b C)
       
    73   then show ?case using rlen_le_size[of C] by simp
       
    74 qed auto
       
    75 
       
    76 text \<open>Now with explicit positions:\<close>
       
    77 
       
    78 datatype dir = L | R
       
    79 
       
    80 type_synonym "pos" = "dir list"
       
    81 
       
    82 function (sequential) rotR_poss :: "'a tree \<Rightarrow> pos list" where
       
    83 "rotR_poss (Node (Node A a B) b C) = [] # rotR_poss (Node A a (Node B b C))" |
       
    84 "rotR_poss (Node Leaf a A) = map (Cons R) (rotR_poss A)" |
       
    85 "rotR_poss Leaf = []"
       
    86 by pat_completeness auto
       
    87 
       
    88 termination
       
    89 proof
       
    90   let ?R = "measure(\<lambda>t. 2*size t - rlen t)"
       
    91   show "wf ?R" by (auto simp add: mlex_prod_def)
       
    92 
       
    93   fix A a B b C
       
    94   show "(Node A a (Node B b C), Node (Node A a B) b C) \<in> ?R"
       
    95     using rlen_le_size[of C] by(simp)
       
    96 
       
    97   fix a A show "(A, Node Leaf a A) \<in> ?R" using rlen_le_size[of A] by(simp)
       
    98 qed
       
    99 
       
   100 fun rotR :: "'a tree \<Rightarrow> 'a tree" where
       
   101 "rotR (Node (Node A a B) b C) = Node A a (Node B b C)"
       
   102 
       
   103 fun rotL :: "'a tree \<Rightarrow> 'a tree" where
       
   104 "rotL (Node A a (Node B b C)) = Node (Node A a B) b C"
       
   105 
       
   106 fun apply_at :: "('a tree \<Rightarrow> 'a tree) \<Rightarrow> pos \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
       
   107   "apply_at f [] t = f t"
       
   108 | "apply_at f (L # ds) (Node l a r) = Node (apply_at f ds l) a r"
       
   109 | "apply_at f (R # ds) (Node l a r) = Node l a (apply_at f ds r)"
       
   110 
       
   111 fun apply_ats :: "('a tree \<Rightarrow> 'a tree) \<Rightarrow> pos list \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
       
   112 "apply_ats _ [] t = t" |
       
   113 "apply_ats f (p#ps) t = apply_ats f ps (apply_at f p t)"
       
   114 
       
   115 lemma apply_ats_append:
       
   116   "apply_ats f (ps\<^sub>1 @ ps\<^sub>2) t = apply_ats f ps\<^sub>2 (apply_ats f ps\<^sub>1 t)"
       
   117 by (induction ps\<^sub>1 arbitrary: t) auto
       
   118 
       
   119 abbreviation "rotRs \<equiv> apply_ats rotR"
       
   120 abbreviation "rotLs \<equiv> apply_ats rotL"
       
   121 
       
   122 lemma apply_ats_map_R: "apply_ats f (map ((#) R) ps) \<langle>l, a, r\<rangle> = Node l a (apply_ats f ps r)"
       
   123 by(induction ps arbitrary: r) auto
       
   124 
       
   125 lemma inorder_rotRs_poss: "inorder (rotRs (rotR_poss t) t) = inorder t"
       
   126 apply(induction t rule: rotR_poss.induct)
       
   127 apply(auto simp: apply_ats_map_R)
       
   128 done
       
   129 
       
   130 lemma is_list_rotRs: "is_list (rotRs (rotR_poss t) t)"
       
   131 apply(induction t rule: rotR_poss.induct)
       
   132 apply(auto simp: apply_ats_map_R)
       
   133 done
       
   134 
       
   135 lemma "is_list (rotRs ps t) \<longrightarrow> length ps \<le> length(rotR_poss t)"
       
   136 quickcheck[expect=counterexample]
       
   137 oops
       
   138 
       
   139 lemma length_rotRs_poss: "length (rotR_poss t) = size t - rlen t"
       
   140 proof(induction t rule: rotR_poss.induct)
       
   141   case (1 A a B b C)
       
   142   then show ?case using rlen_le_size[of C] by simp
       
   143 qed auto
       
   144 
       
   145 lemma is_list_inorder_same:
       
   146   "is_list t1 \<Longrightarrow> is_list t2 \<Longrightarrow> inorder t1 = inorder t2 \<Longrightarrow> t1 = t2"
       
   147 proof(induction t1 arbitrary: t2)
       
   148   case Leaf
       
   149   then show ?case by simp
       
   150 next
       
   151   case Node
       
   152   then show ?case by (cases t2) simp_all
       
   153 qed
       
   154 
       
   155 lemma rot_id: "rotLs (rev (rotR_poss t)) (rotRs (rotR_poss t) t) = t"
       
   156 apply(induction t rule: rotR_poss.induct)
       
   157 apply(auto simp: apply_ats_map_R rev_map apply_ats_append)
       
   158 done
       
   159 
       
   160 corollary tree_to_tree_rotations: assumes "inorder t1 = inorder t2"
       
   161 shows "rotLs (rev (rotR_poss t2)) (rotRs (rotR_poss t1) t1) = t2"
       
   162 proof -
       
   163   have "rotRs (rotR_poss t1) t1 = rotRs (rotR_poss t2) t2" (is "?L = ?R")
       
   164     by (simp add: assms inorder_rotRs_poss is_list_inorder_same is_list_rotRs)
       
   165   hence "rotLs (rev (rotR_poss t2)) ?L = rotLs (rev (rotR_poss t2)) ?R"
       
   166     by simp
       
   167   also have "\<dots> = t2" by(rule rot_id)
       
   168   finally show ?thesis .
       
   169 qed
       
   170 
       
   171 lemma size_rlen_better_ub: "size t - rlen t \<le> size t - 1"
       
   172 by (cases t) auto
       
   173 
       
   174 end