diff -r f13b82281715 -r 806be481aa57 src/HOL/Data_Structures/Array_Braun.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Data_Structures/Array_Braun.thy Wed Oct 17 21:00:53 2018 +0200 @@ -0,0 +1,309 @@ +section "Arrays via Braun Trees" + +theory Array_Braun +imports + Array_Specs + Braun_Tree +begin + +subsection "Array" + +fun lookup1 :: "'a tree \ nat \ 'a" where +"lookup1 (Node l x r) n = (if n=1 then x else lookup1 (if even n then l else r) (n div 2))" + +fun update1 :: "nat \ 'a \ 'a tree \ 'a tree" where +"update1 n x Leaf = Node Leaf x Leaf" | +"update1 n x (Node l a r) = + (if n=1 then Node l x r else + if even n then Node (update1 (n div 2) x l) a r + else Node l a (update1 (n div 2) x r))" + +fun adds :: "'a list \ nat \ 'a tree \ 'a tree" where +"adds [] n t = t" | +"adds (x#xs) n t = adds xs (n+1) (update1 (n+1) x t)" + +fun list :: "'a tree \ 'a list" where +"list Leaf = []" | +"list (Node l x r) = x # splice (list l) (list r)" + + +subsubsection "Functional Correctness" + +lemma size_list: "size(list t) = size t" +by(induction t)(auto) + +lemma minus1_div2: "(n - Suc 0) div 2 = (if odd n then n div 2 else n div 2 - 1)" +by auto arith + +lemma nth_splice: "\ n < size xs + size ys; size ys \ size xs; size xs \ size ys + 1 \ + \ splice xs ys ! n = (if even n then xs else ys) ! (n div 2)" +apply(induction xs ys arbitrary: n rule: splice.induct) +apply (auto simp: nth_Cons' minus1_div2) +done + +lemma div2_in_bounds: + "\ braun (Node l x r); n \ {1..size(Node l x r)}; n > 1 \ \ + (odd n \ n div 2 \ {1..size r}) \ (even n \ n div 2 \ {1..size l})" +by auto arith + +declare upt_Suc[simp del] + + +text \@{const lookup1}\ + +lemma nth_list_lookup1: "\braun t; i < size t\ \ list t ! i = lookup1 t (i+1)" +proof(induction t arbitrary: i) + case Leaf thus ?case by simp +next + case Node + thus ?case using div2_in_bounds[OF Node.prems(1), of "i+1"] + by (auto simp: nth_splice minus1_div2 size_list) +qed + +lemma list_eq_map_lookup1: "braun t \ list t = map (lookup1 t) [1..@{const update1}\ + +lemma size_update1: "\ braun t; n \ {1.. size t} \ \ size(update1 n x t) = size t" +proof(induction t arbitrary: n) + case Leaf thus ?case by simp +next + case Node thus ?case using div2_in_bounds[OF Node.prems] by simp +qed + +lemma braun_update1: "\braun t; n \ {1.. size t} \ \ braun(update1 n x t)" +proof(induction t arbitrary: n) + case Leaf thus ?case by simp +next + case Node thus ?case + using div2_in_bounds[OF Node.prems] by (simp add: size_update1) +qed + +lemma lookup1_update1: "\ braun t; n \ {1.. size t} \ \ + lookup1 (update1 n x t) m = (if n=m then x else lookup1 t m)" +proof(induction t arbitrary: m n) + case Leaf + then show ?case by simp +next + have aux: "\ odd n; odd m \ \ n div 2 = (m::nat) div 2 \ m=n" for m n + using odd_two_times_div_two_succ by fastforce + case Node + thus ?case using div2_in_bounds[OF Node.prems] by (auto simp: aux) +qed + +lemma list_update1: "\ braun t; n \ {1.. size t} \ \ list(update1 n x t) = (list t)[n-1 := x]" +by(auto simp add: list_eq_map_lookup1 list_eq_iff_nth_eq lookup1_update1 size_update1 braun_update1) + +text \A second proof of @{thm list_update1}:\ + +lemma diff1_eq_iff: "n > 0 \ n - Suc 0 = m \ n = m+1" +by arith + +lemma list_update_splice: + "\ n < size xs + size ys; size ys \ size xs; size xs \ size ys + 1 \ \ + (splice xs ys) [n := x] = + (if even n then splice (xs[n div 2 := x]) ys else splice xs (ys[n div 2 := x]))" +by(induction xs ys arbitrary: n rule: splice.induct) (auto split: nat.split) + +lemma list_update2: "\ braun t; n \ {1.. size t} \ \ list(update1 n x t) = (list t)[n-1 := x]" +proof(induction t arbitrary: n) + case Leaf thus ?case by simp +next + case (Node l a r) thus ?case using div2_in_bounds[OF Node.prems] + by(auto simp: list_update_splice diff1_eq_iff size_list split: nat.split) +qed + + +text \@{const adds}\ + +lemma splice_last: shows + "size ys \ size xs \ splice (xs @ [x]) ys = splice xs ys @ [x]" +and "size ys+1 \ size xs \ splice xs (ys @ [y]) = splice xs ys @ [y]" +by(induction xs ys arbitrary: x y rule: splice.induct) (auto) + +lemma list_add_hi: "braun t \ list(update1 (Suc(size t)) x t) = list t @ [x]" +by(induction t)(auto simp: splice_last size_list) + +lemma size_add_hi: "braun t \ m = size t \ size(update1 (Suc m) x t) = size t + 1" +by(induction t arbitrary: m)(auto) + +lemma braun_add_hi: "braun t \ braun(update1 (Suc(size t)) x t)" +by(induction t)(auto simp: size_add_hi) + +lemma size_braun_adds: + "\ braun t; size t = n \ \ size(adds xs n t) = size t + length xs \ braun (adds xs n t)" +by(induction xs arbitrary: t n)(auto simp: braun_add_hi size_add_hi) + +lemma list_adds: "\ braun t; size t = n \ \ list(adds xs n t) = list t @ xs" +by(induction xs arbitrary: t n)(auto simp: size_braun_adds list_add_hi size_add_hi braun_add_hi) + + +subsubsection "Array Implementation" + +interpretation A: Array +where lookup = "\(t,l) n. lookup1 t (n+1)" +and update = "\n x (t,l). (update1 (n+1) x t, l)" +and len = "\(t,l). l" +and array = "\xs. (adds xs 0 Leaf, length xs)" +and invar = "\(t,l). braun t \ l = size t" +and list = "\(t,l). list t" +proof (standard, goal_cases) + case 1 thus ?case by (simp add: nth_list_lookup1 split: prod.splits) +next + case 2 thus ?case by (simp add: list_update1 split: prod.splits) +next + case 3 thus ?case by (simp add: size_list split: prod.splits) +next + case 4 thus ?case by (simp add: list_adds) +next + case 5 thus ?case by (simp add: braun_update1 size_update1 split: prod.splits) +next + case 6 thus ?case by (simp add: size_braun_adds split: prod.splits) +qed + + +subsection "Flexible Array" + +fun add_lo where +"add_lo x Leaf = Node Leaf x Leaf" | +"add_lo x (Node l a r) = Node (add_lo a r) x l" + +fun merge where +"merge Leaf r = r" | +"merge (Node l a r) rr = Node rr a (merge l r)" + +fun del_lo where +"del_lo Leaf = Leaf" | +"del_lo (Node l a r) = merge l r" + +fun del_hi :: "nat \ 'a tree \ 'a tree" where +"del_hi n Leaf = Leaf" | +"del_hi n (Node l x r) = + (if n = 1 then Leaf + else if even n + then Node (del_hi (n div 2) l) x r + else Node l x (del_hi (n div 2) r))" + + +subsubsection "Functional Correctness" + + +text \@{const add_lo}\ + +lemma list_add_lo: "braun t \ list (add_lo a t) = a # list t" +by(induction t arbitrary: a) auto + +lemma braun_add_lo: "braun t \ braun(add_lo x t)" +by(induction t arbitrary: x) (auto simp add: list_add_lo simp flip: size_list) + + +text \@{const del_lo}\ + +lemma list_merge: "braun (Node l x r) \ list(merge l r) = splice (list l) (list r)" +by (induction l r rule: merge.induct) auto + +lemma braun_merge: "braun (Node l x r) \ braun(merge l r)" +by (induction l r rule: merge.induct)(auto simp add: list_merge simp flip: size_list) + +lemma list_del_lo: "braun t \ list(del_lo t) = tl (list t)" +by (cases t) (simp_all add: list_merge) + +lemma braun_del_lo: "braun t \ braun(del_lo t)" +by (cases t) (simp_all add: braun_merge) + + +text \@{const del_hi}\ + +lemma list_Nil_iff: "list t = [] \ t = Leaf" +by(cases t) simp_all + +lemma butlast_splice: "butlast (splice xs ys) = + (if size xs > size ys then splice (butlast xs) ys else splice xs (butlast ys))" +by(induction xs ys rule: splice.induct) (auto) + +lemma list_del_hi: "braun t \ size t = st \ list(del_hi st t) = butlast(list t)" +apply(induction t arbitrary: st) +by(auto simp: list_Nil_iff size_list butlast_splice) + +lemma braun_del_hi: "braun t \ size t = st \ braun(del_hi st t)" +apply(induction t arbitrary: st) +by(auto simp: list_del_hi simp flip: size_list) + + +subsubsection "Flexible Array Implementation" + +interpretation AF: Array_Flex +where lookup = "\(t,l) n. lookup1 t (n+1)" +and update = "\n x (t,l). (update1 (n+1) x t, l)" +and len = "\(t,l). l" +and array = "\xs. (adds xs 0 Leaf, length xs)" +and invar = "\(t,l). braun t \ l = size t" +and list = "\(t,l). list t" +and add_lo = "\x (t,l). (add_lo x t, l+1)" +and del_lo = "\(t,l). (del_lo t, l-1)" +and add_hi = "\x (t,l). (update1 (Suc l) x t, l+1)" +and del_hi = "\(t,l). (del_hi l t, l-1)" +proof (standard, goal_cases) + case 1 thus ?case by (simp add: list_add_lo split: prod.splits) +next + case 2 thus ?case by (simp add: list_del_lo split: prod.splits) +next + case 3 thus ?case by (simp add: list_add_hi braun_add_hi split: prod.splits) +next + case 4 thus ?case by (simp add: list_del_hi split: prod.splits) +next + case 5 thus ?case by (simp add: braun_add_lo list_add_lo flip: size_list split: prod.splits) +next + case 6 thus ?case by (simp add: braun_del_lo list_del_lo flip: size_list split: prod.splits) +next + case 7 thus ?case by (simp add: size_add_hi braun_add_hi split: prod.splits) +next + case 8 thus ?case by (simp add: braun_del_hi list_del_hi flip: size_list split: prod.splits) +qed + + +subsection "Faster" + +fun braun_of_rec :: "'a \ nat \ 'a tree" where +"braun_of_rec x n = (if n=0 then Leaf + else let m = (n-1) div 2 + in if odd n then Node (braun_of_rec x m) x (braun_of_rec x m) + else Node (braun_of_rec x (m + 1)) x (braun_of_rec x m))" + +fun braun2_of :: "'a \ nat \ 'a tree * 'a tree" where +"braun2_of x n = (if n = 0 then (Leaf, Node Leaf x Leaf) + else let (s,t) = braun2_of x ((n-1) div 2) + in if odd n then (Node s x s, Node t x s) else (Node t x s, Node t x t))" + +definition braun_of :: "'a \ nat \ 'a tree" where +"braun_of x n = fst (braun2_of x n)" + +declare braun2_of.simps [simp del] + +lemma braun2_of_size_braun: "braun2_of x n = (s,t) \ size s = n \ size t = n+1 \ braun s \ braun t" +proof(induction x n arbitrary: s t rule: braun2_of.induct) + case (1 x n) + then show ?case + by (auto simp: braun2_of.simps[of x n] split: prod.splits if_splits) presburger+ +qed + +lemma braun2_of_replicate: + "braun2_of x n = (s,t) \ list s = replicate n x \ list t = replicate (n+1) x" +proof(induction x n arbitrary: s t rule: braun2_of.induct) + case (1 x n) + have "x # replicate m x = replicate (m+1) x" for m by simp + with 1 show ?case + apply (auto simp: braun2_of.simps[of x n] replicate.simps(2)[of 0 x] + simp del: replicate.simps(2) split: prod.splits if_splits) + by presburger+ +qed + +corollary braun_braun_of: "braun(braun_of x n)" +unfolding braun_of_def by (metis eq_fst_iff braun2_of_size_braun) + +corollary list_braun_of: "list(braun_of x n) = replicate n x" +unfolding braun_of_def by (metis eq_fst_iff braun2_of_replicate) + +end \ No newline at end of file