| 72099 |      1 | (* Author: Tobias Nipkow *)
 | 
|  |      2 | 
 | 
|  |      3 | section \<open>2-3 Tree from List\<close>
 | 
|  |      4 | 
 | 
|  |      5 | theory Tree23_of_List
 | 
| 79490 |      6 | imports
 | 
|  |      7 |   Tree23
 | 
|  |      8 |   Define_Time_Function
 | 
| 72099 |      9 | begin
 | 
|  |     10 | 
 | 
|  |     11 | text \<open>Linear-time bottom up conversion of a list of items into a complete 2-3 tree
 | 
|  |     12 | whose inorder traversal yields the list of items.\<close>
 | 
|  |     13 | 
 | 
|  |     14 | 
 | 
|  |     15 | subsection "Code"
 | 
|  |     16 | 
 | 
|  |     17 | text \<open>Nonempty lists of 2-3 trees alternating with items, starting and ending with a 2-3 tree:\<close>
 | 
|  |     18 | 
 | 
|  |     19 | datatype 'a tree23s = T "'a tree23" | TTs "'a tree23" "'a" "'a tree23s"
 | 
|  |     20 | 
 | 
|  |     21 | abbreviation "not_T ts == (\<forall>t. ts \<noteq> T t)"
 | 
|  |     22 | 
 | 
|  |     23 | fun len :: "'a tree23s \<Rightarrow> nat" where
 | 
|  |     24 | "len (T _) = 1" |
 | 
|  |     25 | "len (TTs _ _ ts) = len ts + 1"
 | 
|  |     26 | 
 | 
| 72100 |     27 | fun trees :: "'a tree23s \<Rightarrow> 'a tree23 set" where
 | 
|  |     28 | "trees (T t) = {t}" |
 | 
|  |     29 | "trees (TTs t a ts) = {t} \<union> trees ts"
 | 
| 72099 |     30 | 
 | 
|  |     31 | text \<open>Join pairs of adjacent trees:\<close>
 | 
|  |     32 | 
 | 
|  |     33 | fun join_adj :: "'a tree23s \<Rightarrow> 'a tree23s" where
 | 
|  |     34 | "join_adj (TTs t1 a (T t2)) = T(Node2 t1 a t2)" |
 | 
|  |     35 | "join_adj (TTs t1 a (TTs t2 b (T t3))) = T(Node3 t1 a t2 b t3)" |
 | 
| 72121 |     36 | "join_adj (TTs t1 a (TTs t2 b ts)) = TTs (Node2 t1 a t2) b (join_adj ts)"
 | 
| 72099 |     37 | 
 | 
|  |     38 | text \<open>Towards termination of \<open>join_all\<close>:\<close>
 | 
|  |     39 | 
 | 
|  |     40 | lemma len_ge2:
 | 
|  |     41 |   "not_T ts \<Longrightarrow> len ts \<ge> 2"
 | 
|  |     42 | by(cases ts rule: join_adj.cases) auto
 | 
|  |     43 | 
 | 
|  |     44 | lemma [measure_function]: "is_measure len"
 | 
|  |     45 | by(rule is_measure_trivial)
 | 
|  |     46 | 
 | 
|  |     47 | lemma len_join_adj_div2:
 | 
|  |     48 |   "not_T ts \<Longrightarrow> len(join_adj ts) \<le> len ts div 2"
 | 
|  |     49 | by(induction ts rule: join_adj.induct) auto
 | 
|  |     50 | 
 | 
|  |     51 | lemma len_join_adj1: "not_T ts \<Longrightarrow> len(join_adj ts) < len ts"
 | 
|  |     52 | using len_join_adj_div2[of ts] len_ge2[of ts] by simp
 | 
|  |     53 | 
 | 
|  |     54 | corollary len_join_adj2[termination_simp]: "len(join_adj (TTs t a ts)) \<le> len ts"
 | 
|  |     55 | using len_join_adj1[of "TTs t a ts"] by simp
 | 
|  |     56 | 
 | 
|  |     57 | fun join_all :: "'a tree23s \<Rightarrow> 'a tree23" where
 | 
|  |     58 | "join_all (T t) = t" |
 | 
| 72121 |     59 | "join_all ts = join_all (join_adj ts)"
 | 
| 72099 |     60 | 
 | 
|  |     61 | fun leaves :: "'a list \<Rightarrow> 'a tree23s" where
 | 
|  |     62 | "leaves [] = T Leaf" |
 | 
|  |     63 | "leaves (a # as) = TTs Leaf a (leaves as)"
 | 
|  |     64 | 
 | 
|  |     65 | definition tree23_of_list :: "'a list \<Rightarrow> 'a tree23" where
 | 
|  |     66 | "tree23_of_list as = join_all(leaves as)"
 | 
|  |     67 | 
 | 
|  |     68 | 
 | 
|  |     69 | subsection \<open>Functional correctness\<close>
 | 
|  |     70 | 
 | 
|  |     71 | subsubsection \<open>\<open>inorder\<close>:\<close>
 | 
|  |     72 | 
 | 
|  |     73 | fun inorder2 :: "'a tree23s \<Rightarrow> 'a list" where
 | 
|  |     74 | "inorder2 (T t) = inorder t" |
 | 
|  |     75 | "inorder2 (TTs t a ts) = inorder t @ a # inorder2 ts"
 | 
|  |     76 | 
 | 
|  |     77 | lemma inorder2_join_adj: "not_T ts \<Longrightarrow> inorder2(join_adj ts) = inorder2 ts"
 | 
|  |     78 | by (induction ts rule: join_adj.induct) auto
 | 
|  |     79 | 
 | 
|  |     80 | lemma inorder_join_all: "inorder (join_all ts) = inorder2 ts"
 | 
| 72121 |     81 | proof (induction ts rule: join_all.induct)
 | 
|  |     82 |   case 1 thus ?case by simp
 | 
|  |     83 | next
 | 
|  |     84 |   case (2 t a ts)
 | 
|  |     85 |   thus ?case using inorder2_join_adj[of "TTs t a ts"]
 | 
|  |     86 |     by (simp add: le_imp_less_Suc)
 | 
| 72099 |     87 | qed
 | 
|  |     88 | 
 | 
|  |     89 | lemma inorder2_leaves: "inorder2(leaves as) = as"
 | 
|  |     90 | by(induction as) auto
 | 
|  |     91 | 
 | 
| 72122 |     92 | lemma inorder: "inorder(tree23_of_list as) = as"
 | 
| 72099 |     93 | by(simp add: tree23_of_list_def inorder_join_all inorder2_leaves)
 | 
|  |     94 | 
 | 
|  |     95 | subsubsection \<open>Completeness:\<close>
 | 
|  |     96 | 
 | 
|  |     97 | lemma complete_join_adj:
 | 
| 72100 |     98 |   "\<forall>t \<in> trees ts. complete t \<and> height t = n \<Longrightarrow> not_T ts \<Longrightarrow>
 | 
|  |     99 |    \<forall>t \<in> trees (join_adj ts). complete t \<and> height t = Suc n"
 | 
| 72099 |    100 | by (induction ts rule: join_adj.induct) auto
 | 
|  |    101 | 
 | 
|  |    102 | lemma complete_join_all:
 | 
| 72100 |    103 |  "\<forall>t \<in> trees ts. complete t \<and> height t = n \<Longrightarrow> complete (join_all ts)"
 | 
| 72121 |    104 | proof (induction ts arbitrary: n rule: join_all.induct)
 | 
|  |    105 |   case 1 thus ?case by simp
 | 
|  |    106 | next
 | 
|  |    107 |   case (2 t a ts)
 | 
|  |    108 |   thus ?case
 | 
|  |    109 |     apply simp using complete_join_adj[of "TTs t a ts" n, simplified] by blast
 | 
| 72099 |    110 | qed
 | 
|  |    111 | 
 | 
| 72100 |    112 | lemma complete_leaves: "t \<in> trees (leaves as) \<Longrightarrow> complete t \<and> height t = 0"
 | 
| 72099 |    113 | by (induction as) auto
 | 
|  |    114 | 
 | 
|  |    115 | corollary complete: "complete(tree23_of_list as)"
 | 
|  |    116 | by(simp add: tree23_of_list_def complete_leaves complete_join_all[of _ 0])
 | 
|  |    117 | 
 | 
|  |    118 | 
 | 
|  |    119 | subsection "Linear running time"
 | 
|  |    120 | 
 | 
| 79969 |    121 | time_fun join_adj
 | 
|  |    122 | time_fun join_all
 | 
|  |    123 | time_fun leaves
 | 
|  |    124 | time_fun tree23_of_list
 | 
| 72099 |    125 | 
 | 
| 72543 |    126 | lemma T_join_adj: "not_T ts \<Longrightarrow> T_join_adj ts \<le> len ts div 2"
 | 
|  |    127 | by(induction ts rule: T_join_adj.induct) auto
 | 
| 72099 |    128 | 
 | 
| 72543 |    129 | lemma len_ge_1: "len ts \<ge> 1"
 | 
|  |    130 | by(cases ts) auto
 | 
| 72099 |    131 | 
 | 
| 72543 |    132 | lemma T_join_all: "T_join_all ts \<le> 2 * len ts"
 | 
| 72121 |    133 | proof(induction ts rule: join_all.induct)
 | 
|  |    134 |   case 1 thus ?case by simp
 | 
|  |    135 | next
 | 
|  |    136 |   case (2 t a ts)
 | 
|  |    137 |   let ?ts = "TTs t a ts"
 | 
| 72543 |    138 |   have "T_join_all ?ts = T_join_adj ?ts + T_join_all (join_adj ?ts) + 1"
 | 
| 72121 |    139 |     by simp
 | 
| 72543 |    140 |   also have "\<dots> \<le> len ?ts div 2 + T_join_all (join_adj ?ts) + 1"
 | 
|  |    141 |     using T_join_adj[of ?ts] by simp
 | 
|  |    142 |   also have "\<dots> \<le> len ?ts div 2 + 2 * len (join_adj ?ts) + 1"
 | 
| 72121 |    143 |     using "2.IH" by simp
 | 
| 72543 |    144 |   also have "\<dots> \<le> len ?ts div 2 + 2 * (len ?ts div 2) + 1"
 | 
| 72121 |    145 |     using len_join_adj_div2[of ?ts] by simp
 | 
| 72543 |    146 |   also have "\<dots> \<le> 2 * len ?ts" using len_ge_1[of ?ts] by linarith
 | 
| 72121 |    147 |   finally show ?case .
 | 
| 72099 |    148 | qed
 | 
|  |    149 | 
 | 
| 72543 |    150 | lemma T_leaves: "T_leaves as = length as + 1"
 | 
| 72099 |    151 | by(induction as) auto
 | 
|  |    152 | 
 | 
|  |    153 | lemma len_leaves: "len(leaves as) = length as + 1"
 | 
|  |    154 | by(induction as) auto
 | 
|  |    155 | 
 | 
| 79138 |    156 | lemma T_tree23_of_list: "T_tree23_of_list as \<le> 3*(length as) + 3"
 | 
| 79490 |    157 | using T_join_all[of "leaves as"] by(simp add: T_leaves len_leaves)
 | 
| 72099 |    158 | 
 | 
|  |    159 | end
 |