# HG changeset patch # User nipkow # Date 1569334227 -7200 # Node ID fd9614c98dd677626454823cea1510d053cbb1e9 # Parent b3b84b71e398acb6f4890b47dca776b0c4ea96fa simplified proofs diff -r b3b84b71e398 -r fd9614c98dd6 src/HOL/Data_Structures/Balance.thy --- a/src/HOL/Data_Structures/Balance.thy Mon Sep 23 17:15:44 2019 +0200 +++ b/src/HOL/Data_Structures/Balance.thy Tue Sep 24 16:10:27 2019 +0200 @@ -40,7 +40,7 @@ lemma bal_inorder: "\ n \ length xs; bal n xs = (t,zs) \ - \ inorder t = take n xs \ zs = drop n xs" + \ xs = inorder t @ zs \ size t = n" proof(induction n xs arbitrary: t zs rule: bal.induct) case (1 n xs) show ?case proof cases @@ -53,24 +53,18 @@ b2: "bal ?m' (tl ys) = (r,zs)" and t: "t = \l, hd ys, r\" by(auto simp: Let_def bal_simps split: prod.splits) - have IH1: "inorder l = take ?m xs \ ys = drop ?m xs" + have IH1: "xs = inorder l @ ys \ size l = ?m" using b1 "1.prems"(1) by(intro "1.IH"(1)) auto - have IH2: "inorder r = take ?m' (tl ys) \ zs = drop ?m' (tl ys)" + have IH2: "tl ys = inorder r @ zs \ size r = ?m'" using b1 b2 IH1 "1.prems"(1) by(intro "1.IH"(2)) auto - have "drop (n div 2) xs \ []" using "1.prems"(1) by simp - hence "hd (drop ?m xs) # take ?m' (tl (drop ?m xs)) = take (?m' + 1) (drop ?m xs)" - by (metis Suc_eq_plus1 take_Suc) - hence *: "inorder t = take n xs" using t IH1 IH2 - using take_add[of ?m "?m'+1" xs] by(simp) - have "n - n div 2 + n div 2 = n" by simp - hence "zs = drop n xs" using IH1 IH2 by (simp add: drop_Suc[symmetric]) - thus ?thesis using * by blast + show ?thesis using t IH1 IH2 "1.prems"(1) hd_Cons_tl[of ys] by fastforce qed qed corollary inorder_bal_list[simp]: "n \ length xs \ inorder(bal_list n xs) = take n xs" -unfolding bal_list_def by (metis bal_inorder eq_fst_iff) +unfolding bal_list_def +by (metis (mono_tags) prod.collapse[of "bal n xs"] append_eq_conv_conj bal_inorder length_inorder) corollary inorder_balance_list[simp]: "inorder(balance_list xs) = xs" by(simp add: balance_list_def) @@ -83,18 +77,18 @@ by(simp add: balance_tree_def inorder_bal_tree) -text\The size lemmas below do not require the precondition @{prop"n \ length xs"} +text\The length/size lemmas below do not require the precondition @{prop"n \ length xs"} (or @{prop"n \ size t"}) that they come with. They could take advantage of the fact that @{term "bal xs n"} yields a result even if @{prop "n > length xs"}. In that case the result will contain one or more occurrences of @{term "hd []"}. However, this is counter-intuitive and does not reflect the execution in an eager functional language.\ -lemma size_bal: "\ n \ length xs; bal n xs = (t,zs) \ \ size t = n \ length zs = length xs - n" -by (metis bal_inorder length_drop length_inorder length_take min.absorb2) +lemma bal_length: "\ n \ length xs; bal n xs = (t,zs) \ \ length zs = length xs - n" +using bal_inorder by fastforce corollary size_bal_list[simp]: "n \ length xs \ size(bal_list n xs) = n" -unfolding bal_list_def by (metis prod.collapse size_bal) +unfolding bal_list_def using bal_inorder prod.exhaust_sel by blast corollary size_balance_list[simp]: "size(balance_list xs) = length xs" by (simp add: balance_list_def) @@ -105,10 +99,6 @@ corollary size_balance_tree[simp]: "size(balance_tree t) = size t" by(simp add: balance_tree_def) -lemma pre_rec2: "\ n \ length xs; bal (n div 2) xs = (l, ys) \ - \ (n - 1 - n div 2) \ length(tl ys)" -using size_bal[of "n div 2" xs l ys] by simp - lemma min_height_bal: "\ n \ length xs; bal n xs = (t,zs) \ \ min_height t = nat(\log 2 (n + 1)\)" proof(induction n xs arbitrary: t zs rule: bal.induct) @@ -118,23 +108,24 @@ assume "n = 0" thus ?thesis using "1.prems"(2) by (simp add: bal_simps) next assume [arith]: "n \ 0" + let ?m = "n div 2" let ?m' = "n - 1 - ?m" from "1.prems" obtain l r ys where - b1: "bal (n div 2) xs = (l,ys)" and - b2: "bal (n - 1 - n div 2) (tl ys) = (r,zs)" and + b1: "bal ?m xs = (l,ys)" and + b2: "bal ?m' (tl ys) = (r,zs)" and t: "t = \l, hd ys, r\" by(auto simp: bal_simps Let_def split: prod.splits) - let ?log1 = "nat (floor(log 2 (n div 2 + 1)))" - let ?log2 = "nat (floor(log 2 (n - 1 - n div 2 + 1)))" - have IH1: "min_height l = ?log1" using "1.IH"(1) b1 "1.prems"(1) by simp - have IH2: "min_height r = ?log2" - using "1.prems"(1) size_bal[OF _ b1] size_bal[OF _ b2] b1 b2 by(intro "1.IH"(2)) auto + let ?hl = "nat (floor(log 2 (?m + 1)))" + let ?hr = "nat (floor(log 2 (?m' + 1)))" + have IH1: "min_height l = ?hl" using "1.IH"(1) b1 "1.prems"(1) by simp + have IH2: "min_height r = ?hr" + using "1.prems"(1) bal_length[OF _ b1] b1 b2 by(intro "1.IH"(2)) auto have "(n+1) div 2 \ 1" by arith hence 0: "log 2 ((n+1) div 2) \ 0" by simp - have "n - 1 - n div 2 + 1 \ n div 2 + 1" by arith - hence le: "?log2 \ ?log1" by(simp add: nat_mono floor_mono) - have "min_height t = min ?log1 ?log2 + 1" by (simp add: t IH1 IH2) - also have "\ = ?log2 + 1" using le by (simp add: min_absorb2) - also have "n - 1 - n div 2 + 1 = (n+1) div 2" by linarith + have "?m' \ ?m" by arith + hence le: "?hr \ ?hl" by(simp add: nat_mono floor_mono) + have "min_height t = min ?hl ?hr + 1" by (simp add: t IH1 IH2) + also have "\ = ?hr + 1" using le by (simp add: min_absorb2) + also have "?m' + 1 = (n+1) div 2" by linarith also have "nat (floor(log 2 ((n+1) div 2))) + 1 = nat (floor(log 2 ((n+1) div 2) + 1))" using 0 by linarith @@ -153,25 +144,24 @@ using "1.prems" by (simp add: bal_simps) next assume [arith]: "n \ 0" + let ?m = "n div 2" let ?m' = "n - 1 - ?m" from "1.prems" obtain l r ys where - b1: "bal (n div 2) xs = (l,ys)" and - b2: "bal (n - 1 - n div 2) (tl ys) = (r,zs)" and + b1: "bal ?m xs = (l,ys)" and + b2: "bal ?m' (tl ys) = (r,zs)" and t: "t = \l, hd ys, r\" by(auto simp: bal_simps Let_def split: prod.splits) - let ?log1 = "nat \log 2 (n div 2 + 1)\" - let ?log2 = "nat \log 2 (n - 1 - n div 2 + 1)\" - have 1: "n div 2 \ length xs" using "1.prems"(1) by linarith - have 2: "n - 1 - n div 2 \ length (tl ys)" using "1.prems"(1) size_bal[OF 1 b1] by simp - have IH1: "height l = ?log1" using "1.IH"(1) b1 "1.prems"(1) by simp - have IH2: "height r = ?log2" - using b1 b2 size_bal[OF _ b1] size_bal[OF _ b2] "1.prems"(1) by(intro "1.IH"(2)) auto - have 0: "log 2 (n div 2 + 1) \ 0" by simp - have "n - 1 - n div 2 + 1 \ n div 2 + 1" by arith - hence le: "?log2 \ ?log1" + let ?hl = "nat \log 2 (?m + 1)\" + let ?hr = "nat \log 2 (?m' + 1)\" + have IH1: "height l = ?hl" using "1.IH"(1) b1 "1.prems"(1) by simp + have IH2: "height r = ?hr" + using b1 b2 bal_length[OF _ b1] "1.prems"(1) by(intro "1.IH"(2)) auto + have 0: "log 2 (?m + 1) \ 0" by simp + have "?m' \ ?m" by arith + hence le: "?hr \ ?hl" by(simp add: nat_mono ceiling_mono del: nat_ceiling_le_eq) - have "height t = max ?log1 ?log2 + 1" by (simp add: t IH1 IH2) - also have "\ = ?log1 + 1" using le by (simp add: max_absorb1) - also have "\ = nat \log 2 (n div 2 + 1) + 1\" using 0 by linarith + have "height t = max ?hl ?hr + 1" by (simp add: t IH1 IH2) + also have "\ = ?hl + 1" using le by (simp add: max_absorb1) + also have "\ = nat \log 2 (?m + 1) + 1\" using 0 by linarith also have "\ = nat \log 2 (n + 1)\" using ceiling_log2_div2[of "n+1"] by (simp) finally show ?thesis . @@ -229,8 +219,8 @@ by(auto simp add: bal_simps Let_def split: prod.splits) have l: "wbalanced l" using "1.IH"(1)[OF \n\0\ refl _ rec1] "1.prems"(1) by linarith have "wbalanced r" - using rec1 rec2 pre_rec2[OF "1.prems"(1) rec1] by(intro "1.IH"(2)) auto - with l t size_bal[OF _ rec1] size_bal[OF _ rec2] "1.prems"(1) + using rec1 rec2 bal_length[OF _ rec1] "1.prems"(1) by(intro "1.IH"(2)) auto + with l t bal_length[OF _ rec1] "1.prems"(1) bal_inorder[OF _ rec1] bal_inorder[OF _ rec2] show ?thesis by auto qed qed