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