# HG changeset patch # User Thomas Sewell # Date 1553612581 -3600 # Node ID 3afa3b25b5e75e726f1dc7d897ac4a3c75d1af43 # Parent 4ce5ce3a612bf3d5ec6ef6c82d189346fcd7eef7 Tweak Braun tree list_fast_rec recursion. A minor adjustment simplifies the termination argument slightly. diff -r 4ce5ce3a612b -r 3afa3b25b5e7 src/HOL/Data_Structures/Array_Braun.thy --- a/src/HOL/Data_Structures/Array_Braun.thy Mon Mar 25 21:46:37 2019 +0100 +++ b/src/HOL/Data_Structures/Array_Braun.thy Tue Mar 26 16:03:01 2019 +0100 @@ -516,22 +516,29 @@ text \The code and the proof are originally due to Thomas Sewell (except running time).\ function list_fast_rec :: "'a tree list \ 'a list" where -"list_fast_rec ts = (if ts = [] then [] else - let us = filter (\t. t \ Leaf) ts - in map value us @ list_fast_rec (map left us @ map right us))" +"list_fast_rec ts = (let us = filter (\t. t \ Leaf) ts in + if us = [] then [] else + map value us @ list_fast_rec (map left us @ map right us))" by (pat_completeness, auto) -lemma list_fast_rec_term: "\ ts \ []; us = filter (\t. t \ \\) ts \ \ - (map left us @ map right us, ts) \ measure (sum_list \ map (\t. 2 * size t + 1))" -apply (clarsimp simp: sum_list_addf[symmetric] sum_list_map_filter') -apply (rule sum_list_strict_mono; simp) -apply (case_tac x; simp) -done +lemma list_fast_rec_term1: "ts \ [] \ Leaf \ set ts \ + sum_list (map (size o left) ts) + sum_list (map (size o right) ts) < sum_list (map size ts)" + apply (clarsimp simp: sum_list_addf[symmetric] sum_list_map_filter') + apply (rule sum_list_strict_mono; clarsimp?) + apply (case_tac x; simp) + done + +lemma list_fast_rec_term: "us \ [] \ us = filter (\t. t \ \\) ts \ + sum_list (map (size o left) us) + sum_list (map (size o right) us) < sum_list (map size ts)" + apply (rule order_less_le_trans, rule list_fast_rec_term1, simp_all) + apply (rule sum_list_filter_le_nat) + done termination -apply (relation "measure (sum_list o map (\t. 2 * size t + 1))") - apply simp -using list_fast_rec_term by auto + apply (relation "measure (sum_list o map size)") + apply simp + apply (simp add: list_fast_rec_term) + done declare list_fast_rec.simps[simp del] @@ -539,19 +546,19 @@ "list_fast t = list_fast_rec [t]" function t_list_fast_rec :: "'a tree list \ nat" where -"t_list_fast_rec ts = (if ts = [] then 0 else - let us = filter (\t. t \ Leaf) ts - in length ts + 5 * length us + t_list_fast_rec (map left us @ map right us))" +"t_list_fast_rec ts = (let us = filter (\t. t \ Leaf) ts + in if us = [] then 0 else + length ts + 5 * length us + t_list_fast_rec (map left us @ map right us))" by (pat_completeness, auto) termination -apply (relation "measure (sum_list o map (\t. 2 * size t + 1))") - apply simp -using list_fast_rec_term by auto + apply (relation "measure (sum_list o map size)") + apply simp + apply (simp add: list_fast_rec_term) + done declare t_list_fast_rec.simps[simp del] - paragraph "Functional Correctness" lemma list_fast_rec_all_Leaf: @@ -588,7 +595,7 @@ apply(auto simp: nth_append braun_list_not_Nil take_nths_eq_single braun_list_Nil hd_drop_conv_nth) done thus ?thesis - by (clarsimp simp: list_fast_rec.simps[of ts] o_def list_fast_rec_all_Leaf) + by (clarsimp simp: list_fast_rec.simps[of ts] o_def list_fast_rec_all_Leaf Let_def) next case False with less.prems(2) have *: @@ -614,7 +621,6 @@ "braun t \ list_fast t = list t" by (simp add: list_fast_def take_nths_00 braun_list_eq list_fast_rec_correct[where k=0]) - paragraph "Running Time Analysis" lemma sum_tree_list_children: "\t \ set ts. t \ Leaf \ @@ -623,21 +629,20 @@ theorem t_list_fast_rec_ub: "t_list_fast_rec ts \ sum_list (map (\t. 7*size t + 1) ts)" -proof (induction ts rule: measure_induct_rule[where f="sum_list o map (\t. 2*size t + 1)"]) +proof (induction ts rule: measure_induct_rule[where f="sum_list o map size"]) case (less ts) + let ?us = "filter (\t. t \ Leaf) ts" show ?case proof cases - assume "ts = []" + assume "?us = []" thus ?thesis using t_list_fast_rec.simps[of ts] by(simp add: Let_def) next - assume "ts \ []" - let ?us = "filter (\t. t \ Leaf) ts" + assume "?us \ []" let ?children = "map left ?us @ map right ?us" have "t_list_fast_rec ts = t_list_fast_rec ?children + 5 * length ?us + length ts" - using \ts \ []\ t_list_fast_rec.simps[of ts] by(simp add: Let_def) + using \?us \ []\ t_list_fast_rec.simps[of ts] by(simp add: Let_def) also have "\ \ (\t\?children. 7 * size t + 1) + 5 * length ?us + length ts" - using less[of "map left ?us @ map right ?us"] - list_fast_rec_term[of ts, OF \ts \ []\] + using less[of "?children"] list_fast_rec_term[of "?us"] \?us \ []\ by (simp) also have "\ = (\t\?children. 7*size t) + 7 * length ?us + length ts" by(simp add: sum_list_Suc o_def)