# HG changeset patch # User nipkow # Date 1569424977 -7200 # Node ID 3fb16bed5d6c6ceb6e3e4129dcd91abc262ad300 # Parent c5232e6fb10b42b791bbccd20acfa1a923d9abf7 replaced new type ('a,'b) tree by old type ('a*'b) tree. diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/AA_Map.thy --- a/src/HOL/Data_Structures/AA_Map.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/AA_Map.thy Wed Sep 25 17:22:57 2019 +0200 @@ -9,21 +9,21 @@ begin fun update :: "'a::linorder \ 'b \ ('a*'b) aa_tree \ ('a*'b) aa_tree" where -"update x y Leaf = Node Leaf (x,y) 1 Leaf" | -"update x y (Node t1 (a,b) lv t2) = +"update x y Leaf = Node Leaf ((x,y), 1) Leaf" | +"update x y (Node t1 ((a,b), lv) t2) = (case cmp x a of - LT \ split (skew (Node (update x y t1) (a,b) lv t2)) | - GT \ split (skew (Node t1 (a,b) lv (update x y t2))) | - EQ \ Node t1 (x,y) lv t2)" + LT \ split (skew (Node (update x y t1) ((a,b), lv) t2)) | + GT \ split (skew (Node t1 ((a,b), lv) (update x y t2))) | + EQ \ Node t1 ((x,y), lv) t2)" fun delete :: "'a::linorder \ ('a*'b) aa_tree \ ('a*'b) aa_tree" where "delete _ Leaf = Leaf" | -"delete x (Node l (a,b) lv r) = +"delete x (Node l ((a,b), lv) r) = (case cmp x a of - LT \ adjust (Node (delete x l) (a,b) lv r) | - GT \ adjust (Node l (a,b) lv (delete x r)) | + LT \ adjust (Node (delete x l) ((a,b), lv) r) | + GT \ adjust (Node l ((a,b), lv) (delete x r)) | EQ \ (if l = Leaf then r - else let (l',ab') = split_max l in adjust (Node l' ab' lv r)))" + else let (l',ab') = split_max l in adjust (Node l' (ab', lv) r)))" subsection "Invariance" @@ -64,20 +64,20 @@ qed simp lemma lvl_update_incr_iff: "(lvl(update a b t) = lvl t + 1) \ - (\l x r. update a b t = Node l x (lvl t + 1) r \ lvl l = lvl r)" + (\l x r. update a b t = Node l (x,lvl t + 1) r \ lvl l = lvl r)" apply(cases t) apply(auto simp add: skew_case split_case split: if_splits) apply(auto split: tree.splits if_splits) done lemma invar_update: "invar t \ invar(update a b t)" -proof(induction t) +proof(induction t rule: tree2_induct) case N: (Node l xy n r) hence il: "invar l" and ir: "invar r" by auto note iil = N.IH(1)[OF il] note iir = N.IH(2)[OF ir] obtain x y where [simp]: "xy = (x,y)" by fastforce - let ?t = "Node l xy n r" + let ?t = "Node l (xy, n) r" have "a < x \ a = x \ x < a" by auto moreover have ?case if "a < x" @@ -87,16 +87,16 @@ by (simp add: skew_invar split_invar del: invar.simps) next case (Incr) - then obtain t1 w t2 where ial[simp]: "update a b l = Node t1 w n t2" + then obtain t1 w t2 where ial[simp]: "update a b l = Node t1 (w, n) t2" using N.prems by (auto simp: lvl_Suc_iff) have l12: "lvl t1 = lvl t2" by (metis Incr(1) ial lvl_update_incr_iff tree.inject) - have "update a b ?t = split(skew(Node (update a b l) xy n r))" + have "update a b ?t = split(skew(Node (update a b l) (xy, n) r))" by(simp add: \a) - also have "skew(Node (update a b l) xy n r) = Node t1 w n (Node t2 xy n r)" + also have "skew(Node (update a b l) (xy, n) r) = Node t1 (w, n) (Node t2 (xy, n) r)" by(simp) also have "invar(split \)" - proof (cases r) + proof (cases r rule: tree2_cases) case Leaf hence "l = Leaf" using N.prems by(auto simp: lvl_0_iff) thus ?thesis using Leaf ial by simp @@ -118,14 +118,14 @@ thus ?case proof assume 0: "n = lvl r" - have "update a b ?t = split(skew(Node l xy n (update a b r)))" + have "update a b ?t = split(skew(Node l (xy, n) (update a b r)))" using \a>x\ by(auto) - also have "skew(Node l xy n (update a b r)) = Node l xy n (update a b r)" + also have "skew(Node l (xy, n) (update a b r)) = Node l (xy, n) (update a b r)" using N.prems by(simp add: skew_case split: tree.split) also have "invar(split \)" proof - from lvl_update_sngl[OF ir sngl_if_invar[OF \invar ?t\ 0], of a b] - obtain t1 p t2 where iar: "update a b r = Node t1 p n t2" + obtain t1 p t2 where iar: "update a b r = Node t1 (p, n) t2" using N.prems 0 by (auto simp: lvl_Suc_iff) from N.prems iar 0 iir show ?thesis by (auto simp: split_case split: tree.splits) @@ -156,13 +156,13 @@ declare invar.simps(2)[simp del] theorem post_delete: "invar t \ post_del t (delete x t)" -proof (induction t) +proof (induction t rule: tree2_induct) case (Node l ab lv r) obtain a b where [simp]: "ab = (a,b)" by fastforce let ?l' = "delete x l" and ?r' = "delete x r" - let ?t = "Node l ab lv r" let ?t' = "delete x ?t" + let ?t = "Node l (ab, lv) r" let ?t' = "delete x ?t" from Node.prems have inv_l: "invar l" and inv_r: "invar r" by (auto) diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/AA_Set.thy --- a/src/HOL/Data_Structures/AA_Set.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/AA_Set.thy Wed Sep 25 17:22:57 2019 +0200 @@ -10,67 +10,67 @@ Cmp begin -type_synonym 'a aa_tree = "('a,nat) tree" +type_synonym 'a aa_tree = "('a*nat) tree" definition empty :: "'a aa_tree" where "empty = Leaf" fun lvl :: "'a aa_tree \ nat" where "lvl Leaf = 0" | -"lvl (Node _ _ lv _) = lv" +"lvl (Node _ (_, lv) _) = lv" fun invar :: "'a aa_tree \ bool" where "invar Leaf = True" | -"invar (Node l a h r) = +"invar (Node l (a, h) r) = (invar l \ invar r \ - h = lvl l + 1 \ (h = lvl r + 1 \ (\lr b rr. r = Node lr b h rr \ h = lvl rr + 1)))" + h = lvl l + 1 \ (h = lvl r + 1 \ (\lr b rr. r = Node lr (b,h) rr \ h = lvl rr + 1)))" fun skew :: "'a aa_tree \ 'a aa_tree" where -"skew (Node (Node t1 b lvb t2) a lva t3) = - (if lva = lvb then Node t1 b lvb (Node t2 a lva t3) else Node (Node t1 b lvb t2) a lva t3)" | +"skew (Node (Node t1 (b, lvb) t2) (a, lva) t3) = + (if lva = lvb then Node t1 (b, lvb) (Node t2 (a, lva) t3) else Node (Node t1 (b, lvb) t2) (a, lva) t3)" | "skew t = t" fun split :: "'a aa_tree \ 'a aa_tree" where -"split (Node t1 a lva (Node t2 b lvb (Node t3 c lvc t4))) = +"split (Node t1 (a, lva) (Node t2 (b, lvb) (Node t3 (c, lvc) t4))) = (if lva = lvb \ lvb = lvc \ \\lva = lvc\ suffices\ - then Node (Node t1 a lva t2) b (lva+1) (Node t3 c lva t4) - else Node t1 a lva (Node t2 b lvb (Node t3 c lvc t4)))" | + then Node (Node t1 (a,lva) t2) (b,lva+1) (Node t3 (c, lva) t4) + else Node t1 (a,lva) (Node t2 (b,lvb) (Node t3 (c,lvc) t4)))" | "split t = t" hide_const (open) insert fun insert :: "'a::linorder \ 'a aa_tree \ 'a aa_tree" where -"insert x Leaf = Node Leaf x 1 Leaf" | -"insert x (Node t1 a lv t2) = +"insert x Leaf = Node Leaf (x, 1) Leaf" | +"insert x (Node t1 (a,lv) t2) = (case cmp x a of - LT \ split (skew (Node (insert x t1) a lv t2)) | - GT \ split (skew (Node t1 a lv (insert x t2))) | - EQ \ Node t1 x lv t2)" + LT \ split (skew (Node (insert x t1) (a,lv) t2)) | + GT \ split (skew (Node t1 (a,lv) (insert x t2))) | + EQ \ Node t1 (x, lv) t2)" fun sngl :: "'a aa_tree \ bool" where "sngl Leaf = False" | -"sngl (Node _ _ _ Leaf) = True" | -"sngl (Node _ _ lva (Node _ _ lvb _)) = (lva > lvb)" +"sngl (Node _ _ Leaf) = True" | +"sngl (Node _ (_, lva) (Node _ (_, lvb) _)) = (lva > lvb)" definition adjust :: "'a aa_tree \ 'a aa_tree" where "adjust t = (case t of - Node l x lv r \ + Node l (x,lv) r \ (if lvl l >= lv-1 \ lvl r >= lv-1 then t else - if lvl r < lv-1 \ sngl l then skew (Node l x (lv-1) r) else + if lvl r < lv-1 \ sngl l then skew (Node l (x,lv-1) r) else if lvl r < lv-1 then case l of - Node t1 a lva (Node t2 b lvb t3) - \ Node (Node t1 a lva t2) b (lvb+1) (Node t3 x (lv-1) r) + Node t1 (a,lva) (Node t2 (b,lvb) t3) + \ Node (Node t1 (a,lva) t2) (b,lvb+1) (Node t3 (x,lv-1) r) else - if lvl r < lv then split (Node l x (lv-1) r) + if lvl r < lv then split (Node l (x,lv-1) r) else case r of - Node t1 b lvb t4 \ + Node t1 (b,lvb) t4 \ (case t1 of - Node t2 a lva t3 - \ Node (Node l x (lv-1) t2) a (lva+1) - (split (Node t3 b (if sngl t1 then lva else lva+1) t4)))))" + Node t2 (a,lva) t3 + \ Node (Node l (x,lv-1) t2) (a,lva+1) + (split (Node t3 (b, if sngl t1 then lva else lva+1) t4)))))" text\In the paper, the last case of \<^const>\adjust\ is expressed with the help of an incorrect auxiliary function \texttt{nlvl}. @@ -81,20 +81,20 @@ is not restored.\ fun split_max :: "'a aa_tree \ 'a aa_tree * 'a" where -"split_max (Node l a lv Leaf) = (l,a)" | -"split_max (Node l a lv r) = (let (r',b) = split_max r in (adjust(Node l a lv r'), b))" +"split_max (Node l (a,lv) Leaf) = (l,a)" | +"split_max (Node l (a,lv) r) = (let (r',b) = split_max r in (adjust(Node l (a,lv) r'), b))" fun delete :: "'a::linorder \ 'a aa_tree \ 'a aa_tree" where "delete _ Leaf = Leaf" | -"delete x (Node l a lv r) = +"delete x (Node l (a,lv) r) = (case cmp x a of - LT \ adjust (Node (delete x l) a lv r) | - GT \ adjust (Node l a lv (delete x r)) | + LT \ adjust (Node (delete x l) (a,lv) r) | + GT \ adjust (Node l (a,lv) (delete x r)) | EQ \ (if l = Leaf then r - else let (l',b) = split_max l in adjust (Node l' b lv r)))" + else let (l',b) = split_max l in adjust (Node l' (b,lv) r)))" fun pre_adjust where -"pre_adjust (Node l a lv r) = (invar l \ invar r \ +"pre_adjust (Node l (a,lv) r) = (invar l \ invar r \ ((lv = lvl l + 1 \ (lv = lvl r + 1 \ lv = lvl r + 2 \ lv = lvl r \ sngl r)) \ (lv = lvl l + 2 \ (lv = lvl r + 1 \ lv = lvl r \ sngl r))))" @@ -103,23 +103,23 @@ subsection "Auxiliary Proofs" lemma split_case: "split t = (case t of - Node t1 x lvx (Node t2 y lvy (Node t3 z lvz t4)) \ + Node t1 (x,lvx) (Node t2 (y,lvy) (Node t3 (z,lvz) t4)) \ (if lvx = lvy \ lvy = lvz - then Node (Node t1 x lvx t2) y (lvx+1) (Node t3 z lvx t4) + then Node (Node t1 (x,lvx) t2) (y,lvx+1) (Node t3 (z,lvx) t4) else t) | t \ t)" by(auto split: tree.split) lemma skew_case: "skew t = (case t of - Node (Node t1 y lvy t2) x lvx t3 \ - (if lvx = lvy then Node t1 y lvx (Node t2 x lvx t3) else t) + Node (Node t1 (y,lvy) t2) (x,lvx) t3 \ + (if lvx = lvy then Node t1 (y, lvx) (Node t2 (x,lvx) t3) else t) | t \ t)" by(auto split: tree.split) lemma lvl_0_iff: "invar t \ lvl t = 0 \ t = Leaf" by(cases t) auto -lemma lvl_Suc_iff: "lvl t = Suc n \ (\ l a r. t = Node l a (Suc n) r)" +lemma lvl_Suc_iff: "lvl t = Suc n \ (\ l a r. t = Node l (a,Suc n) r)" by(cases t) auto lemma lvl_skew: "lvl (skew t) = lvl t" @@ -128,16 +128,16 @@ lemma lvl_split: "lvl (split t) = lvl t \ lvl (split t) = lvl t + 1 \ sngl (split t)" by(cases t rule: split.cases) auto -lemma invar_2Nodes:"invar (Node l x lv (Node rl rx rlv rr)) = - (invar l \ invar \rl, rx, rlv, rr\ \ lv = Suc (lvl l) \ +lemma invar_2Nodes:"invar (Node l (x,lv) (Node rl (rx, rlv) rr)) = + (invar l \ invar \rl, (rx, rlv), rr\ \ lv = Suc (lvl l) \ (lv = Suc rlv \ rlv = lv \ lv = Suc (lvl rr)))" by simp lemma invar_NodeLeaf[simp]: - "invar (Node l x lv Leaf) = (invar l \ lv = Suc (lvl l) \ lv = Suc 0)" + "invar (Node l (x,lv) Leaf) = (invar l \ lv = Suc (lvl l) \ lv = Suc 0)" by simp -lemma sngl_if_invar: "invar (Node l a n r) \ n = lvl r \ sngl r" +lemma sngl_if_invar: "invar (Node l (a, n) r) \ n = lvl r \ sngl r" by(cases r rule: sngl.cases) clarsimp+ @@ -167,7 +167,8 @@ thus ?thesis using 2 by (auto simp add: skew_case split_case split: tree.splits) next case GT - thus ?thesis using 2 proof (cases t1) + thus ?thesis using 2 + proof (cases t1 rule: tree2_cases) case Node thus ?thesis using 2 GT apply (auto simp add: skew_case split_case split: tree.splits) @@ -183,32 +184,32 @@ by(cases t rule: split.cases) clarsimp+ lemma invar_NodeL: - "\ invar(Node l x n r); invar l'; lvl l' = lvl l \ \ invar(Node l' x n r)" + "\ invar(Node l (x, n) r); invar l'; lvl l' = lvl l \ \ invar(Node l' (x, n) r)" by(auto) lemma invar_NodeR: - "\ invar(Node l x n r); n = lvl r + 1; invar r'; lvl r' = lvl r \ \ invar(Node l x n r')" + "\ invar(Node l (x, n) r); n = lvl r + 1; invar r'; lvl r' = lvl r \ \ invar(Node l (x, n) r')" by(auto) lemma invar_NodeR2: - "\ invar(Node l x n r); sngl r'; n = lvl r + 1; invar r'; lvl r' = n \ \ invar(Node l x n r')" + "\ invar(Node l (x, n) r); sngl r'; n = lvl r + 1; invar r'; lvl r' = n \ \ invar(Node l (x, n) r')" by(cases r' rule: sngl.cases) clarsimp+ lemma lvl_insert_incr_iff: "(lvl(insert a t) = lvl t + 1) \ - (\l x r. insert a t = Node l x (lvl t + 1) r \ lvl l = lvl r)" -apply(cases t) + (\l x r. insert a t = Node l (x, lvl t + 1) r \ lvl l = lvl r)" +apply(cases t rule: tree2_cases) apply(auto simp add: skew_case split_case split: if_splits) apply(auto split: tree.splits if_splits) done lemma invar_insert: "invar t \ invar(insert a t)" -proof(induction t) +proof(induction t rule: tree2_induct) case N: (Node l x n r) hence il: "invar l" and ir: "invar r" by auto note iil = N.IH(1)[OF il] note iir = N.IH(2)[OF ir] - let ?t = "Node l x n r" + let ?t = "Node l (x, n) r" have "a < x \ a = x \ x < a" by auto moreover have ?case if "a < x" @@ -218,16 +219,16 @@ by (simp add: skew_invar split_invar del: invar.simps) next case (Incr) - then obtain t1 w t2 where ial[simp]: "insert a l = Node t1 w n t2" + then obtain t1 w t2 where ial[simp]: "insert a l = Node t1 (w, n) t2" using N.prems by (auto simp: lvl_Suc_iff) have l12: "lvl t1 = lvl t2" by (metis Incr(1) ial lvl_insert_incr_iff tree.inject) - have "insert a ?t = split(skew(Node (insert a l) x n r))" + have "insert a ?t = split(skew(Node (insert a l) (x,n) r))" by(simp add: \a) - also have "skew(Node (insert a l) x n r) = Node t1 w n (Node t2 x n r)" + also have "skew(Node (insert a l) (x,n) r) = Node t1 (w,n) (Node t2 (x,n) r)" by(simp) also have "invar(split \)" - proof (cases r) + proof (cases r rule: tree2_cases) case Leaf hence "l = Leaf" using N.prems by(auto simp: lvl_0_iff) thus ?thesis using Leaf ial by simp @@ -249,14 +250,14 @@ thus ?case proof assume 0: "n = lvl r" - have "insert a ?t = split(skew(Node l x n (insert a r)))" + have "insert a ?t = split(skew(Node l (x, n) (insert a r)))" using \a>x\ by(auto) - also have "skew(Node l x n (insert a r)) = Node l x n (insert a r)" + also have "skew(Node l (x,n) (insert a r)) = Node l (x,n) (insert a r)" using N.prems by(simp add: skew_case split: tree.split) also have "invar(split \)" proof - from lvl_insert_sngl[OF ir sngl_if_invar[OF \invar ?t\ 0], of a] - obtain t1 y t2 where iar: "insert a r = Node t1 y n t2" + obtain t1 y t2 where iar: "insert a r = Node t1 (y,n) t2" using N.prems 0 by (auto simp: lvl_Suc_iff) from N.prems iar 0 iir show ?thesis by (auto simp: split_case split: tree.splits) @@ -285,21 +286,21 @@ subsubsection "Proofs for delete" -lemma invarL: "ASSUMPTION(invar \l, a, lv, r\) \ invar l" +lemma invarL: "ASSUMPTION(invar \l, (a, lv), r\) \ invar l" by(simp add: ASSUMPTION_def) -lemma invarR: "ASSUMPTION(invar \lv, l, a, r\) \ invar r" +lemma invarR: "ASSUMPTION(invar \l, (a,lv), r\) \ invar r" by(simp add: ASSUMPTION_def) lemma sngl_NodeI: - "sngl (Node l a lv r) \ sngl (Node l' a' lv r)" -by(cases r) (simp_all) + "sngl (Node l (a,lv) r) \ sngl (Node l' (a', lv) r)" +by(cases r rule: tree2_cases) (simp_all) declare invarL[simp] invarR[simp] lemma pre_cases: -assumes "pre_adjust (Node l x lv r)" +assumes "pre_adjust (Node l (x,lv) r)" obtains (tSngl) "invar l \ invar r \ lv = Suc (lvl r) \ lvl l = lvl r" | @@ -317,38 +318,39 @@ declare invar.simps(2)[simp del] invar_2Nodes[simp add] lemma invar_adjust: - assumes pre: "pre_adjust (Node l a lv r)" - shows "invar(adjust (Node l a lv r))" + assumes pre: "pre_adjust (Node l (a,lv) r)" + shows "invar(adjust (Node l (a,lv) r))" using pre proof (cases rule: pre_cases) case (tDouble) thus ?thesis unfolding adjust_def by (cases r) (auto simp: invar.simps(2)) next case (rDown) - from rDown obtain llv ll la lr where l: "l = Node ll la llv lr" by (cases l) auto + from rDown obtain llv ll la lr where l: "l = Node ll (la, llv) lr" by (cases l) auto from rDown show ?thesis unfolding adjust_def by (auto simp: l invar.simps(2) split: tree.splits) next case (lDown_tDouble) - from lDown_tDouble obtain rlv rr ra rl where r: "r = Node rl ra rlv rr" by (cases r) auto + from lDown_tDouble obtain rlv rr ra rl where r: "r = Node rl (ra, rlv) rr" by (cases r) auto from lDown_tDouble and r obtain rrlv rrr rra rrl where - rr :"rr = Node rrr rra rrlv rrl" by (cases rr) auto + rr :"rr = Node rrr (rra, rrlv) rrl" by (cases rr) auto from lDown_tDouble show ?thesis unfolding adjust_def r rr - apply (cases rl) apply (auto simp add: invar.simps(2) split!: if_split) + apply (cases rl rule: tree2_cases) apply (auto simp add: invar.simps(2) split!: if_split) using lDown_tDouble by (auto simp: split_case lvl_0_iff elim:lvl.elims split: tree.split) qed (auto simp: split_case invar.simps(2) adjust_def split: tree.splits) lemma lvl_adjust: - assumes "pre_adjust (Node l a lv r)" - shows "lv = lvl (adjust(Node l a lv r)) \ lv = lvl (adjust(Node l a lv r)) + 1" -using assms(1) proof(cases rule: pre_cases) + assumes "pre_adjust (Node l (a,lv) r)" + shows "lv = lvl (adjust(Node l (a,lv) r)) \ lv = lvl (adjust(Node l (a,lv) r)) + 1" +using assms(1) +proof(cases rule: pre_cases) case lDown_tSngl thus ?thesis - using lvl_split[of "\l, a, lvl r, r\"] by (auto simp: adjust_def) + using lvl_split[of "\l, (a, lvl r), r\"] by (auto simp: adjust_def) next case lDown_tDouble thus ?thesis by (auto simp: adjust_def invar.simps(2) split: tree.split) qed (auto simp: adjust_def split: tree.splits) -lemma sngl_adjust: assumes "pre_adjust (Node l a lv r)" - "sngl \l, a, lv, r\" "lv = lvl (adjust \l, a, lv, r\)" - shows "sngl (adjust \l, a, lv, r\)" +lemma sngl_adjust: assumes "pre_adjust (Node l (a,lv) r)" + "sngl \l, (a, lv), r\" "lv = lvl (adjust \l, (a, lv), r\)" + shows "sngl (adjust \l, (a, lv), r\)" using assms proof (cases rule: pre_cases) case rDown thus ?thesis using assms(2,3) unfolding adjust_def @@ -361,38 +363,38 @@ (lvl t' = lvl t \ sngl t \ sngl t')" lemma pre_adj_if_postR: - "invar\lv, l, a, r\ \ post_del r r' \ pre_adjust \lv, l, a, r'\" + "invar\lv, (l, a), r\ \ post_del r r' \ pre_adjust \lv, (l, a), r'\" by(cases "sngl r") (auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims) lemma pre_adj_if_postL: - "invar\l, a, lv, r\ \ post_del l l' \ pre_adjust \l', b, lv, r\" + "invar\l, (a, lv), r\ \ post_del l l' \ pre_adjust \l', (b, lv), r\" by(cases "sngl r") (auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims) lemma post_del_adjL: - "\ invar\l, a, lv, r\; pre_adjust \l', b, lv, r\ \ - \ post_del \l, a, lv, r\ (adjust \l', b, lv, r\)" + "\ invar\l, (a, lv), r\; pre_adjust \l', (b, lv), r\ \ + \ post_del \l, (a, lv), r\ (adjust \l', (b, lv), r\)" unfolding post_del_def by (metis invar_adjust lvl_adjust sngl_NodeI sngl_adjust lvl.simps(2)) lemma post_del_adjR: -assumes "invar\lv, l, a, r\" "pre_adjust \lv, l, a, r'\" "post_del r r'" -shows "post_del \lv, l, a, r\ (adjust \lv, l, a, r'\)" +assumes "invar\l, (a,lv), r\" "pre_adjust \l, (a,lv), r'\" "post_del r r'" +shows "post_del \l, (a,lv), r\ (adjust \l, (a,lv), r'\)" proof(unfold post_del_def, safe del: disjCI) - let ?t = "\lv, l, a, r\" - let ?t' = "adjust \lv, l, a, r'\" + let ?t = "\l, (a,lv), r\" + let ?t' = "adjust \l, (a,lv), r'\" show "invar ?t'" by(rule invar_adjust[OF assms(2)]) show "lvl ?t' = lvl ?t \ lvl ?t' + 1 = lvl ?t" using lvl_adjust[OF assms(2)] by auto show "sngl ?t'" if as: "lvl ?t' = lvl ?t" "sngl ?t" proof - - have s: "sngl \lv, l, a, r'\" - proof(cases r') + have s: "sngl \l, (a,lv), r'\" + proof(cases r' rule: tree2_cases) case Leaf thus ?thesis by simp next case Node thus ?thesis using as(2) assms(1,3) - by (cases r) (auto simp: post_del_def) + by (cases r rule: tree2_cases) (auto simp: post_del_def) qed show ?thesis using as(1) sngl_adjust[OF assms(2) s] by simp qed @@ -403,22 +405,22 @@ theorem post_split_max: "\ invar t; (t', x) = split_max t; t \ Leaf \ \ post_del t t'" proof (induction t arbitrary: t' rule: split_max.induct) - case (2 lv l a lvr rl ra rr) - let ?r = "\lvr, rl, ra, rr\" - let ?t = "\lv, l, a, ?r\" + case (2 l a lv rl bl rr) + let ?r = "\rl, bl, rr\" + let ?t = "\l, (a, lv), ?r\" from "2.prems"(2) obtain r' where r': "(r', x) = split_max ?r" - and [simp]: "t' = adjust \lv, l, a, r'\" by auto + and [simp]: "t' = adjust \l, (a, lv), r'\" by auto from "2.IH"[OF _ r'] \invar ?t\ have post: "post_del ?r r'" by simp note preR = pre_adj_if_postR[OF \invar ?t\ post] show ?case by (simp add: post_del_adjR[OF "2.prems"(1) preR post]) qed (auto simp: post_del_def) theorem post_delete: "invar t \ post_del t (delete x t)" -proof (induction t) +proof (induction t rule: tree2_induct) case (Node l a lv r) let ?l' = "delete x l" and ?r' = "delete x r" - let ?t = "Node l a lv r" let ?t' = "delete x ?t" + let ?t = "Node l (a,lv) r" let ?t' = "delete x ?t" from Node.prems have inv_l: "invar l" and inv_r: "invar r" by (auto) diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/AVL_Map.thy --- a/src/HOL/Data_Structures/AVL_Map.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/AVL_Map.thy Wed Sep 25 17:22:57 2019 +0200 @@ -9,16 +9,16 @@ begin fun update :: "'a::linorder \ 'b \ ('a*'b) avl_tree \ ('a*'b) avl_tree" where -"update x y Leaf = Node Leaf (x,y) 1 Leaf" | -"update x y (Node l (a,b) h r) = (case cmp x a of - EQ \ Node l (x,y) h r | +"update x y Leaf = Node Leaf ((x,y), 1) Leaf" | +"update x y (Node l ((a,b), h) r) = (case cmp x a of + EQ \ Node l ((x,y), h) r | LT \ balL (update x y l) (a,b) r | GT \ balR l (a,b) (update x y r))" fun delete :: "'a::linorder \ ('a*'b) avl_tree \ ('a*'b) avl_tree" where "delete _ Leaf = Leaf" | -"delete x (Node l (a,b) h r) = (case cmp x a of - EQ \ del_root (Node l (a,b) h r) | +"delete x (Node l ((a,b), h) r) = (case cmp x a of + EQ \ del_root (Node l ((a,b), h) r) | LT \ balR (delete x l) (a,b) r | GT \ balL l (a,b) (delete x r))" @@ -114,7 +114,7 @@ assumes "avl t" shows "avl(delete x t)" and "height t = (height (delete x t)) \ height t = height (delete x t) + 1" using assms -proof (induct t) +proof (induct t rule: tree2_induct) case (Node l n h r) obtain a b where [simp]: "n = (a,b)" by fastforce case 1 @@ -134,8 +134,8 @@ show ?case proof(cases "x = a") case True - with 1 have "height (Node l n h r) = height(del_root (Node l n h r)) - \ height (Node l n h r) = height(del_root (Node l n h r)) + 1" + with 1 have "height (Node l (n, h) r) = height(del_root (Node l (n, h) r)) + \ height (Node l (n, h) r) = height(del_root (Node l (n, h) r)) + 1" by (subst height_del_root,simp_all) with True show ?thesis by simp next diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/AVL_Set.thy --- a/src/HOL/Data_Structures/AVL_Set.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/AVL_Set.thy Wed Sep 25 17:22:57 2019 +0200 @@ -12,7 +12,7 @@ "HOL-Number_Theory.Fib" begin -type_synonym 'a avl_tree = "('a,nat) tree" +type_synonym 'a avl_tree = "('a*nat) tree" definition empty :: "'a avl_tree" where "empty = Leaf" @@ -21,25 +21,25 @@ fun avl :: "'a avl_tree \ bool" where "avl Leaf = True" | -"avl (Node l a h r) = +"avl (Node l (a,h) r) = ((height l = height r \ height l = height r + 1 \ height r = height l + 1) \ h = max (height l) (height r) + 1 \ avl l \ avl r)" fun ht :: "'a avl_tree \ nat" where "ht Leaf = 0" | -"ht (Node l a h r) = h" +"ht (Node l (a,h) r) = h" definition node :: "'a avl_tree \ 'a \ 'a avl_tree \ 'a avl_tree" where -"node l a r = Node l a (max (ht l) (ht r) + 1) r" +"node l a r = Node l (a, max (ht l) (ht r) + 1) r" definition balL :: "'a avl_tree \ 'a \ 'a avl_tree \ 'a avl_tree" where "balL l a r = (if ht l = ht r + 2 then case l of - Node bl b _ br \ + Node bl (b, _) br \ if ht bl < ht br then case br of - Node cl c _ cr \ node (node bl b cl) c (node cr a r) + Node cl (c, _) cr \ node (node bl b cl) c (node cr a r) else node bl b (node br a r) else node l a r)" @@ -47,38 +47,38 @@ "balR l a r = (if ht r = ht l + 2 then case r of - Node bl b _ br \ + Node bl (b, _) br \ if ht bl > ht br then case bl of - Node cl c _ cr \ node (node l a cl) c (node cr b br) + Node cl (c, _) cr \ node (node l a cl) c (node cr b br) else node (node l a bl) b br else node l a r)" fun insert :: "'a::linorder \ 'a avl_tree \ 'a avl_tree" where -"insert x Leaf = Node Leaf x 1 Leaf" | -"insert x (Node l a h r) = (case cmp x a of - EQ \ Node l a h r | +"insert x Leaf = Node Leaf (x, 1) Leaf" | +"insert x (Node l (a, h) r) = (case cmp x a of + EQ \ Node l (a, h) r | LT \ balL (insert x l) a r | GT \ balR l a (insert x r))" fun split_max :: "'a avl_tree \ 'a avl_tree * 'a" where -"split_max (Node l a _ r) = +"split_max (Node l (a, _) r) = (if r = Leaf then (l,a) else let (r',a') = split_max r in (balL l a r', a'))" lemmas split_max_induct = split_max.induct[case_names Node Leaf] fun del_root :: "'a avl_tree \ 'a avl_tree" where -"del_root (Node Leaf a h r) = r" | -"del_root (Node l a h Leaf) = l" | -"del_root (Node l a h r) = (let (l', a') = split_max l in balR l' a' r)" +"del_root (Node Leaf (a,h) r) = r" | +"del_root (Node l (a,h) Leaf) = l" | +"del_root (Node l (a,h) r) = (let (l', a') = split_max l in balR l' a' r)" -lemmas del_root_cases = del_root.cases[case_names Leaf_t Node_Leaf Node_Node] +lemmas del_root_cases = del_root.cases[split_format(complete), case_names Leaf_t Node_Leaf Node_Node] fun delete :: "'a::linorder \ 'a avl_tree \ 'a avl_tree" where "delete _ Leaf = Leaf" | -"delete x (Node l a h r) = +"delete x (Node l (a, h) r) = (case cmp x a of - EQ \ del_root (Node l a h r) | + EQ \ del_root (Node l (a, h) r) | LT \ balR (delete x l) a r | GT \ balL l a (delete x r))" @@ -113,8 +113,8 @@ (auto simp: inorder_balL split: if_splits prod.splits tree.split) lemma inorder_del_root: - "inorder (del_root (Node l a h r)) = inorder l @ inorder r" -by(cases "Node l a h r" rule: del_root.cases) + "inorder (del_root (Node l ah r)) = inorder l @ inorder r" +by(cases "Node l ah r" rule: del_root.cases) (auto simp: inorder_balL inorder_balR inorder_split_maxD split: if_splits prod.splits) theorem inorder_delete: @@ -134,7 +134,7 @@ declare Let_def [simp] lemma ht_height[simp]: "avl t \ ht t = height t" -by (cases t) simp_all +by (cases t rule: tree2_cases) simp_all lemma height_balL: "\ height l = height r + 2; avl l; avl r \ \ @@ -171,7 +171,7 @@ assumes "avl l" "avl r" and "height l = height r \ height l = height r + 1 \ height r = height l + 1 \ height l = height r + 2" shows "avl(balL l a r)" -proof(cases l) +proof(cases l rule: tree2_cases) case Leaf with assms show ?thesis by (simp add: node_def balL_def) next @@ -191,7 +191,7 @@ assumes "avl l" and "avl r" and "height l = height r \ height l = height r + 1 \ height r = height l + 1 \ height r = height l + 2" shows "avl(balR l a r)" -proof(cases r) +proof(cases r rule: tree2_cases) case Leaf with assms show ?thesis by (simp add: node_def balR_def) next @@ -216,7 +216,7 @@ shows "avl(insert x t)" "(height (insert x t) = height t \ height (insert x t) = height t + 1)" using assms -proof (induction t) +proof (induction t rule: tree2_induct) case (Node l a h r) case 1 show ?case @@ -304,8 +304,8 @@ using assms proof (cases t rule:del_root_cases) case (Node_Node ll ln lh lr n h rl rn rh rr) - let ?l = "Node ll ln lh lr" - let ?r = "Node rl rn rh rr" + let ?l = "Node ll (ln, lh) lr" + let ?r = "Node rl (rn, rh) rr" let ?l' = "fst (split_max ?l)" from \avl t\ and Node_Node have "avl ?r" by simp from \avl t\ and Node_Node have "avl ?l" by simp @@ -324,8 +324,8 @@ using assms proof (cases t rule: del_root_cases) case (Node_Node ll ln lh lr n h rl rn rh rr) - let ?l = "Node ll ln lh lr" - let ?r = "Node rl rn rh rr" + let ?l = "Node ll (ln, lh) lr" + let ?r = "Node rl (rn, rh) rr" let ?l' = "fst (split_max ?l)" let ?t' = "balR ?l' (snd(split_max ?l)) ?r" from \avl t\ and Node_Node have "avl ?r" by simp @@ -356,7 +356,7 @@ assumes "avl t" shows "avl(delete x t)" and "height t = (height (delete x t)) \ height t = height (delete x t) + 1" using assms -proof (induct t) +proof (induct t rule: tree2_induct) case (Node l n h r) case 1 show ?case @@ -375,8 +375,8 @@ show ?case proof(cases "x = n") case True - with 1 have "height (Node l n h r) = height(del_root (Node l n h r)) - \ height (Node l n h r) = height(del_root (Node l n h r)) + 1" + with 1 have "height (Node l (n,h) r) = height(del_root (Node l (n,h) r)) + \ height (Node l (n,h) r) = height(del_root (Node l (n,h) r)) + 1" by (subst height_del_root,simp_all) with True show ?thesis by simp next @@ -449,7 +449,7 @@ lemma height_invers: "(height t = 0) = (t = Leaf)" - "avl t \ (height t = Suc h) = (\ l a r . t = Node l a (Suc h) r)" + "avl t \ (height t = Suc h) = (\ l a r . t = Node l (a,Suc h) r)" by (induction t) auto text \Any AVL tree of height \h\ has at least \fib (h+2)\ leaves:\ @@ -462,7 +462,7 @@ next case (3 h) from "3.prems" obtain l a r where - [simp]: "t = Node l a (Suc(Suc h)) r" "avl l" "avl r" + [simp]: "t = Node l (a,Suc(Suc h)) r" "avl l" "avl r" and C: " height r = Suc h \ height l = Suc h \ height r = Suc h \ height l = h diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/Isin2.thy --- a/src/HOL/Data_Structures/Isin2.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/Isin2.thy Wed Sep 25 17:22:57 2019 +0200 @@ -9,18 +9,18 @@ Set_Specs begin -fun isin :: "('a::linorder,'b) tree \ 'a \ bool" where +fun isin :: "('a::linorder*'b) tree \ 'a \ bool" where "isin Leaf x = False" | -"isin (Node l a _ r) x = +"isin (Node l (a,_) r) x = (case cmp x a of LT \ isin l x | EQ \ True | GT \ isin r x)" lemma isin_set_inorder: "sorted(inorder t) \ isin t x = (x \ set(inorder t))" -by (induction t) (auto simp: isin_simps) +by (induction t rule: tree2_induct) (auto simp: isin_simps) lemma isin_set_tree: "bst t \ isin t x \ x \ set_tree t" -by(induction t) auto +by(induction t rule: tree2_induct) auto end diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/Leftist_Heap.thy --- a/src/HOL/Data_Structures/Leftist_Heap.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/Leftist_Heap.thy Wed Sep 25 17:22:57 2019 +0200 @@ -10,30 +10,30 @@ Complex_Main begin -fun mset_tree :: "('a,'b) tree \ 'a multiset" where +fun mset_tree :: "('a*'b) tree \ 'a multiset" where "mset_tree Leaf = {#}" | -"mset_tree (Node l a _ r) = {#a#} + mset_tree l + mset_tree r" +"mset_tree (Node l (a, _) r) = {#a#} + mset_tree l + mset_tree r" -type_synonym 'a lheap = "('a,nat)tree" +type_synonym 'a lheap = "('a*nat)tree" fun rank :: "'a lheap \ nat" where "rank Leaf = 0" | -"rank (Node _ _ _ r) = rank r + 1" +"rank (Node _ _ r) = rank r + 1" fun rk :: "'a lheap \ nat" where "rk Leaf = 0" | -"rk (Node _ _ n _) = n" +"rk (Node _ (_, n) _) = n" text\The invariants:\ -fun (in linorder) heap :: "('a,'b) tree \ bool" where +fun (in linorder) heap :: "('a*'b) tree \ bool" where "heap Leaf = True" | -"heap (Node l m _ r) = +"heap (Node l (m, _) r) = (heap l \ heap r \ (\x \ set_tree l \ set_tree r. m \ x))" fun ltree :: "'a lheap \ bool" where "ltree Leaf = True" | -"ltree (Node l a n r) = +"ltree (Node l (a, n) r) = (n = rank r + 1 \ rank l \ rank r \ ltree l & ltree r)" definition empty :: "'a lheap" where @@ -42,10 +42,10 @@ definition node :: "'a lheap \ 'a \ 'a lheap \ 'a lheap" where "node l a r = (let rl = rk l; rr = rk r - in if rl \ rr then Node l a (rr+1) r else Node r a (rl+1) l)" + in if rl \ rr then Node l (a,rr+1) r else Node r (a,rl+1) l)" fun get_min :: "'a lheap \ 'a" where -"get_min(Node l a n r) = a" +"get_min(Node l (a, n) r) = a" text \For function \merge\:\ unbundle pattern_aliases @@ -53,7 +53,7 @@ fun merge :: "'a::ord lheap \ 'a lheap \ 'a lheap" where "merge Leaf t = t" | "merge t Leaf = t" | -"merge (Node l1 a1 n1 r1 =: t1) (Node l2 a2 n2 r2 =: t2) = +"merge (Node l1 (a1, n1) r1 =: t1) (Node l2 (a2, n2) r2 =: t2) = (if a1 \ a2 then node l1 a1 (merge r1 t2) else node l2 a2 (merge t1 r2))" @@ -63,18 +63,18 @@ lemma merge_code: "merge t1 t2 = (case (t1,t2) of (Leaf, _) \ t2 | (_, Leaf) \ t1 | - (Node l1 a1 n1 r1, Node l2 a2 n2 r2) \ + (Node l1 (a1, n1) r1, Node l2 (a2, n2) r2) \ if a1 \ a2 then node l1 a1 (merge r1 t2) else node l2 a2 (merge t1 r2))" by(induction t1 t2 rule: merge.induct) (simp_all split: tree.split) hide_const (open) insert definition insert :: "'a::ord \ 'a lheap \ 'a lheap" where -"insert x t = merge (Node Leaf x 1 Leaf) t" +"insert x t = merge (Node Leaf (x,1) Leaf) t" fun del_min :: "'a::ord lheap \ 'a lheap" where "del_min Leaf = Leaf" | -"del_min (Node l x n r) = merge l r" +"del_min (Node l _ r) = merge l r" subsection "Lemmas" @@ -177,17 +177,17 @@ subsection "Complexity" lemma pow2_rank_size1: "ltree t \ 2 ^ rank t \ size1 t" -proof(induction t) +proof(induction t rule: tree2_induct) case Leaf show ?case by simp next case (Node l a n r) hence "rank r \ rank l" by simp hence *: "(2::nat) ^ rank r \ 2 ^ rank l" by simp - have "(2::nat) ^ rank \l, a, n, r\ = 2 ^ rank r + 2 ^ rank r" + have "(2::nat) ^ rank \l, (a, n), r\ = 2 ^ rank r + 2 ^ rank r" by(simp add: mult_2) also have "\ \ size1 l + size1 r" using Node * by (simp del: power_increasing_iff) - also have "\ = size1 \l, a, n, r\" by simp + also have "\ = size1 \l, (a, n), r\" by simp finally show ?case . qed @@ -196,16 +196,16 @@ fun t_merge :: "'a::ord lheap \ 'a lheap \ nat" where "t_merge Leaf t = 1" | "t_merge t Leaf = 1" | -"t_merge (Node l1 a1 n1 r1 =: t1) (Node l2 a2 n2 r2 =: t2) = +"t_merge (Node l1 (a1, n1) r1 =: t1) (Node l2 (a2, n2) r2 =: t2) = (if a1 \ a2 then 1 + t_merge r1 t2 else 1 + t_merge t1 r2)" definition t_insert :: "'a::ord \ 'a lheap \ nat" where -"t_insert x t = t_merge (Node Leaf x 1 Leaf) t" +"t_insert x t = t_merge (Node Leaf (x, 1) Leaf) t" fun t_del_min :: "'a::ord lheap \ nat" where "t_del_min Leaf = 1" | -"t_del_min (Node l a n r) = t_merge l r" +"t_del_min (Node l _ r) = t_merge l r" lemma t_merge_rank: "t_merge l r \ rank l + rank r + 1" proof(induction l r rule: merge.induct) @@ -219,7 +219,7 @@ by linarith corollary t_insert_log: "ltree t \ t_insert x t \ log 2 (size1 t) + 2" -using t_merge_log[of "Node Leaf x 1 Leaf" t] +using t_merge_log[of "Node Leaf (x, 1) Leaf" t] by(simp add: t_insert_def split: tree.split) (* FIXME mv ? *) @@ -237,7 +237,7 @@ corollary t_del_min_log: assumes "ltree t" shows "t_del_min t \ 2 * log 2 (size1 t) + 1" -proof(cases t) +proof(cases t rule: tree2_cases) case Leaf thus ?thesis using assms by simp next case [simp]: (Node t1 _ _ t2) diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/Lookup2.thy --- a/src/HOL/Data_Structures/Lookup2.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/Lookup2.thy Wed Sep 25 17:22:57 2019 +0200 @@ -9,13 +9,13 @@ Map_Specs begin -fun lookup :: "('a::linorder * 'b, 'c) tree \ 'a \ 'b option" where +fun lookup :: "(('a::linorder * 'b) * 'c) tree \ 'a \ 'b option" where "lookup Leaf x = None" | -"lookup (Node l (a,b) _ r) x = +"lookup (Node l ((a,b), _) r) x = (case cmp x a of LT \ lookup l x | GT \ lookup r x | EQ \ Some b)" lemma lookup_map_of: "sorted1(inorder t) \ lookup t x = map_of (inorder t) x" -by(induction t) (auto simp: map_of_simps split: option.split) +by(induction t rule: tree2_induct) (auto simp: map_of_simps split: option.split) end diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/RBT.thy --- a/src/HOL/Data_Structures/RBT.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/RBT.thy Wed Sep 25 17:22:57 2019 +0200 @@ -8,10 +8,10 @@ datatype color = Red | Black -type_synonym 'a rbt = "('a,color)tree" +type_synonym 'a rbt = "('a*color)tree" -abbreviation R where "R l a r \ Node l a Red r" -abbreviation B where "B l a r \ Node l a Black r" +abbreviation R where "R l a r \ Node l (a, Red) r" +abbreviation B where "B l a r \ Node l (a, Black) r" fun baliL :: "'a rbt \ 'a \ 'a rbt \ 'a rbt" where "baliL (R (R t1 a t2) b t3) c t4 = R (B t1 a t2) b (B t3 c t4)" | @@ -25,7 +25,7 @@ fun paint :: "color \ 'a rbt \ 'a rbt" where "paint c Leaf = Leaf" | -"paint c (Node l a _ r) = Node l a c r" +"paint c (Node l (a,_) r) = Node l (a,c) r" fun baldL :: "'a rbt \ 'a \ 'a rbt \ 'a rbt" where "baldL (R t1 a t2) b t3 = R (B t1 a t2) b t3" | diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/RBT_Map.thy --- a/src/HOL/Data_Structures/RBT_Map.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/RBT_Map.thy Wed Sep 25 17:22:57 2019 +0200 @@ -24,7 +24,7 @@ fun del :: "'a::linorder \ ('a*'b)rbt \ ('a*'b)rbt" where "del x Leaf = Leaf" | -"del x (Node l (a,b) c r) = (case cmp x a of +"del x (Node l ((a,b), c) r) = (case cmp x a of LT \ if l \ Leaf \ color l = Black then baldL (del x l) (a,b) r else R (del x l) (a,b) r | GT \ if r \ Leaf\ color r = Black diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/RBT_Set.thy --- a/src/HOL/Data_Structures/RBT_Set.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/RBT_Set.thy Wed Sep 25 17:22:57 2019 +0200 @@ -31,11 +31,11 @@ fun color :: "'a rbt \ color" where "color Leaf = Black" | -"color (Node _ _ c _) = c" +"color (Node _ (_, c) _) = c" fun del :: "'a::linorder \ 'a rbt \ 'a rbt" where "del x Leaf = Leaf" | -"del x (Node l a _ r) = +"del x (Node l (a, _) r) = (case cmp x a of LT \ if l \ Leaf \ color l = Black then baldL (del x l) a r else R (del x l) a r | @@ -100,11 +100,11 @@ fun bheight :: "'a rbt \ nat" where "bheight Leaf = 0" | -"bheight (Node l x c r) = (if c = Black then bheight l + 1 else bheight l)" +"bheight (Node l (x, c) r) = (if c = Black then bheight l + 1 else bheight l)" fun invc :: "'a rbt \ bool" where "invc Leaf = True" | -"invc (Node l a c r) = +"invc (Node l (a,c) r) = (invc l \ invc r \ (c = Red \ color l = Black \ color r = Black))" text \Weaker version:\ @@ -113,10 +113,10 @@ fun invh :: "'a rbt \ bool" where "invh Leaf = True" | -"invh (Node l x c r) = (invh l \ invh r \ bheight l = bheight r)" +"invh (Node l (x, c) r) = (invh l \ invh r \ bheight l = bheight r)" lemma invc2I: "invc t \ invc2 t" -by (cases t) simp+ +by (cases t rule: tree2_cases) simp+ definition rbt :: "'a rbt \ bool" where "rbt t = (invc t \ invh t \ color t = Black)" @@ -234,8 +234,8 @@ by (induct l r rule: combine.induct) (auto simp: invc_baldL invc2I split: tree.splits color.splits) -lemma neq_LeafD: "t \ Leaf \ \c l x r. t = Node c l x r" -by(cases t) auto +lemma neq_LeafD: "t \ Leaf \ \l x c r. t = Node l (x,c) r" +by(cases t rule: tree2_cases) auto lemma del_invc_invh: "invh t \ invc t \ invh (del x t) \ (color t = Red \ bheight (del x t) = bheight t \ invc (del x t) \ diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/Set2_Join.thy --- a/src/HOL/Data_Structures/Set2_Join.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/Set2_Join.thy Wed Sep 25 17:22:57 2019 +0200 @@ -20,26 +20,26 @@ and recursion operators on it.\ locale Set2_Join = -fixes join :: "('a::linorder,'b) tree \ 'a \ ('a,'b) tree \ ('a,'b) tree" -fixes inv :: "('a,'b) tree \ bool" +fixes join :: "('a::linorder*'b) tree \ 'a \ ('a*'b) tree \ ('a*'b) tree" +fixes inv :: "('a*'b) tree \ bool" assumes set_join: "set_tree (join l a r) = set_tree l \ {a} \ set_tree r" -assumes bst_join: "bst (Node l a b r) \ bst (join l a r)" +assumes bst_join: "bst (Node l (a, b) r) \ bst (join l a r)" assumes inv_Leaf: "inv \\" assumes inv_join: "\ inv l; inv r \ \ inv (join l a r)" -assumes inv_Node: "\ inv (Node l a b r) \ \ inv l \ inv r" +assumes inv_Node: "\ inv (Node l (a,b) r) \ \ inv l \ inv r" begin declare set_join [simp] subsection "\split_min\" -fun split_min :: "('a,'b) tree \ 'a \ ('a,'b) tree" where -"split_min (Node l a _ r) = +fun split_min :: "('a*'b) tree \ 'a \ ('a*'b) tree" where +"split_min (Node l (a, _) r) = (if l = Leaf then (a,r) else let (m,l') = split_min l in (m, join l' a r))" lemma split_min_set: "\ split_min t = (m,t'); t \ Leaf \ \ m \ set_tree t \ set_tree t = {m} \ set_tree t'" -proof(induction t arbitrary: t') +proof(induction t arbitrary: t' rule: tree2_induct) case Node thus ?case by(auto split: prod.splits if_splits dest: inv_Node) next case Leaf thus ?case by simp @@ -47,7 +47,7 @@ lemma split_min_bst: "\ split_min t = (m,t'); bst t; t \ Leaf \ \ bst t' \ (\x \ set_tree t'. m < x)" -proof(induction t arbitrary: t') +proof(induction t arbitrary: t' rule: tree2_induct) case Node thus ?case by(fastforce simp: split_min_set bst_join split: prod.splits if_splits) next case Leaf thus ?case by simp @@ -55,7 +55,7 @@ lemma split_min_inv: "\ split_min t = (m,t'); inv t; t \ Leaf \ \ inv t'" -proof(induction t arbitrary: t') +proof(induction t arbitrary: t' rule: tree2_induct) case Node thus ?case by(auto simp: inv_join split: prod.splits if_splits dest: inv_Node) next case Leaf thus ?case by simp @@ -64,7 +64,7 @@ subsection "\join2\" -definition join2 :: "('a,'b) tree \ ('a,'b) tree \ ('a,'b) tree" where +definition join2 :: "('a*'b) tree \ ('a*'b) tree \ ('a*'b) tree" where "join2 l r = (if r = Leaf then l else let (m,r') = split_min r in join l m r')" lemma set_join2[simp]: "set_tree (join2 l r) = set_tree l \ set_tree r" @@ -80,9 +80,9 @@ subsection "\split\" -fun split :: "('a,'b)tree \ 'a \ ('a,'b)tree \ bool \ ('a,'b)tree" where +fun split :: "('a*'b)tree \ 'a \ ('a*'b)tree \ bool \ ('a*'b)tree" where "split Leaf k = (Leaf, False, Leaf)" | -"split (Node l a _ r) x = +"split (Node l (a, _) r) x = (case cmp x a of LT \ let (l1,b,l2) = split l x in (l1, b, join l2 a r) | GT \ let (r1,b,r2) = split r x in (join l a r1, b, r2) | @@ -91,14 +91,14 @@ lemma split: "split t x = (l,xin,r) \ bst t \ set_tree l = {a \ set_tree t. a < x} \ set_tree r = {a \ set_tree t. x < a} \ (xin = (x \ set_tree t)) \ bst l \ bst r" -proof(induction t arbitrary: l xin r) +proof(induction t arbitrary: l xin r rule: tree2_induct) case Leaf thus ?case by simp next case Node thus ?case by(force split!: prod.splits if_splits intro!: bst_join) qed lemma split_inv: "split t x = (l,xin,r) \ inv t \ inv l \ inv r" -proof(induction t arbitrary: l xin r) +proof(induction t arbitrary: l xin r rule: tree2_induct) case Leaf thus ?case by simp next case Node @@ -110,7 +110,7 @@ subsection "\insert\" -definition insert :: "'a \ ('a,'b) tree \ ('a,'b) tree" where +definition insert :: "'a \ ('a*'b) tree \ ('a*'b) tree" where "insert x t = (let (l,_,r) = split t x in join l x r)" lemma set_tree_insert: "bst t \ set_tree (insert x t) = {x} \ set_tree t" @@ -125,7 +125,7 @@ subsection "\delete\" -definition delete :: "'a \ ('a,'b) tree \ ('a,'b) tree" where +definition delete :: "'a \ ('a*'b) tree \ ('a*'b) tree" where "delete x t = (let (l,_,r) = split t x in join2 l r)" lemma set_tree_delete: "bst t \ set_tree (delete x t) = set_tree t - {x}" @@ -140,11 +140,11 @@ subsection "\union\" -fun union :: "('a,'b)tree \ ('a,'b)tree \ ('a,'b)tree" where +fun union :: "('a*'b)tree \ ('a*'b)tree \ ('a*'b)tree" where "union t1 t2 = (if t1 = Leaf then t2 else if t2 = Leaf then t1 else - case t1 of Node l1 a _ r1 \ + case t1 of Node l1 (a, _) r1 \ let (l2,_ ,r2) = split t2 a; l' = union l1 l2; r' = union r1 r2 in join l' a r')" @@ -176,11 +176,11 @@ subsection "\inter\" -fun inter :: "('a,'b)tree \ ('a,'b)tree \ ('a,'b)tree" where +fun inter :: "('a*'b)tree \ ('a*'b)tree \ ('a*'b)tree" where "inter t1 t2 = (if t1 = Leaf then Leaf else if t2 = Leaf then Leaf else - case t1 of Node l1 a _ r1 \ + case t1 of Node l1 (a, _) r1 \ let (l2,ain,r2) = split t2 a; l' = inter l1 l2; r' = inter r1 r2 in if ain then join l' a r' else join2 l' r')" @@ -192,7 +192,7 @@ proof(induction t1 t2 rule: inter.induct) case (1 t1 t2) show ?case - proof (cases t1) + proof (cases t1 rule: tree2_cases) case Leaf thus ?thesis by (simp add: inter.simps) next case [simp]: (Node l1 a _ r1) @@ -209,9 +209,9 @@ **: "?L2 \ ?R2 = {}" "a \ ?L2 \ ?R2" "?L1 \ ?R2 = {}" "?L2 \ ?R1 = {}" using split[OF sp] \bst t1\ \bst t2\ by (force, force, force, force, force) have IHl: "set_tree (inter l1 l2) = set_tree l1 \ set_tree l2" - using "1.IH"(1)[OF _ False _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp + using "1.IH"(1)[OF _ False _ _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp have IHr: "set_tree (inter r1 r2) = set_tree r1 \ set_tree r2" - using "1.IH"(2)[OF _ False _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp + using "1.IH"(2)[OF _ False _ _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp have "set_tree t1 \ set_tree t2 = (?L1 \ ?R1 \ {a}) \ (?L2 \ ?R2 \ ?K)" by(simp add: t2) also have "\ = (?L1 \ ?L2) \ (?R1 \ ?R2) \ ?K" @@ -241,11 +241,11 @@ subsection "\diff\" -fun diff :: "('a,'b)tree \ ('a,'b)tree \ ('a,'b)tree" where +fun diff :: "('a*'b)tree \ ('a*'b)tree \ ('a*'b)tree" where "diff t1 t2 = (if t1 = Leaf then Leaf else if t2 = Leaf then t1 else - case t2 of Node l2 a _ r2 \ + case t2 of Node l2 (a, _) r2 \ let (l1,_,r1) = split t1 a; l' = diff l1 l2; r' = diff r1 r2 in join2 l' r')" @@ -257,7 +257,7 @@ proof(induction t1 t2 rule: diff.induct) case (1 t1 t2) show ?case - proof (cases t2) + proof (cases t2 rule: tree2_cases) case Leaf thus ?thesis by (simp add: diff.simps) next case [simp]: (Node l2 a _ r2) @@ -273,9 +273,9 @@ **: "a \ ?L1 \ ?R1" "?L1 \ ?R2 = {}" "?L2 \ ?R1 = {}" using split[OF sp] \bst t1\ \bst t2\ by (force, force, force, force) have IHl: "set_tree (diff l1 l2) = set_tree l1 - set_tree l2" - using "1.IH"(1)[OF False _ _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp + using "1.IH"(1)[OF False _ _ _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp have IHr: "set_tree (diff r1 r2) = set_tree r1 - set_tree r2" - using "1.IH"(2)[OF False _ _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp + using "1.IH"(2)[OF False _ _ _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp have "set_tree t1 - set_tree t2 = (?L1 \ ?R1) - (?L2 \ ?R2 \ {a})" by(simp add: t1) also have "\ = (?L1 - ?L2) \ (?R1 - ?R2)" @@ -340,7 +340,7 @@ end interpretation unbal: Set2_Join -where join = "\l x r. Node l x () r" and inv = "\t. True" +where join = "\l x r. Node l (x, ()) r" and inv = "\t. True" proof (standard, goal_cases) case 1 show ?case by simp next diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/Set2_Join_RBT.thy --- a/src/HOL/Data_Structures/Set2_Join_RBT.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/Set2_Join_RBT.thy Wed Sep 25 17:22:57 2019 +0200 @@ -179,7 +179,7 @@ by(cases "(l,a,r)" rule: baliR.cases) (auto simp: ball_Un) lemma bst_joinL: - "\bst (Node l a n r); bheight l \ bheight r\ + "\bst (Node l (a, n) r); bheight l \ bheight r\ \ bst (joinL l a r)" proof(induction l a r rule: joinL.induct) case (1 l a r) @@ -202,7 +202,7 @@ by(cases t) auto lemma bst_join: - "bst (Node l a n r) \ bst (join l a r)" + "bst (Node l (a, n) r) \ bst (join l a r)" by(auto simp: bst_paint bst_joinL bst_joinR join_def) lemma inv_join: "\ invc l; invh l; invc r; invh r \ \ invc(join l x r) \ invh(join l x r)" diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/Tree2.thy --- a/src/HOL/Data_Structures/Tree2.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/Tree2.thy Wed Sep 25 17:22:57 2019 +0200 @@ -1,40 +1,29 @@ theory Tree2 -imports Main +imports "HOL-Library.Tree" begin -datatype ('a,'b) tree = - Leaf ("\\") | - Node "('a,'b)tree" 'a 'b "('a,'b) tree" ("(1\_,/ _,/ _,/ _\)") - -fun inorder :: "('a,'b)tree \ 'a list" where -"inorder Leaf = []" | -"inorder (Node l a _ r) = inorder l @ a # inorder r" +text \This theory provides the basic infrastructure for the type @{typ \('a * 'b) tree\} +of augmented trees where @{typ 'a} is the key and @{typ 'b} some additional information.\ -fun height :: "('a,'b) tree \ nat" where -"height Leaf = 0" | -"height (Node l a _ r) = max (height l) (height r) + 1" +text \IMPORTANT: Inductions and cases analyses on augmented trees need to use the following +two rules explicitly. They generate nodes of the form @{term "Node l (a,b) r"} +rather than @{term "Node l a r"} for trees of type @{typ "'a tree"}.\ -fun set_tree :: "('a,'b) tree \ 'a set" where -"set_tree Leaf = {}" | -"set_tree (Node l a _ r) = Set.insert a (set_tree l \ set_tree r)" +lemmas tree2_induct = tree.induct[where 'a = "'a * 'b", split_format(complete)] + +lemmas tree2_cases = tree.exhaust[where 'a = "'a * 'b", split_format(complete)] -fun bst :: "('a::linorder,'b) tree \ bool" where -"bst Leaf = True" | -"bst (Node l a _ r) = (bst l \ bst r \ (\x \ set_tree l. x < a) \ (\x \ set_tree r. a < x))" - -fun size1 :: "('a,'b) tree \ nat" where -"size1 \\ = 1" | -"size1 \l, _, _, r\ = size1 l + size1 r" +fun inorder :: "('a*'b)tree \ 'a list" where +"inorder Leaf = []" | +"inorder (Node l (a,_) r) = inorder l @ a # inorder r" -fun complete :: "('a,'b) tree \ bool" where -"complete Leaf = True" | -"complete (Node l _ _ r) = (complete l \ complete r \ height l = height r)" +fun set_tree :: "('a*'b) tree \ 'a set" where +"set_tree Leaf = {}" | +"set_tree (Node l (a,_) r) = Set.insert a (set_tree l \ set_tree r)" -lemma size1_size: "size1 t = size t + 1" -by (induction t) simp_all - -lemma size1_ge0[simp]: "0 < size1 t" -by (simp add: size1_size) +fun bst :: "('a::linorder*'b) tree \ bool" where +"bst Leaf = True" | +"bst (Node l (a, _) r) = (bst l \ bst r \ (\x \ set_tree l. x < a) \ (\x \ set_tree r. a < x))" lemma finite_set_tree[simp]: "finite(set_tree t)" by(induction t) auto diff -r c5232e6fb10b -r 3fb16bed5d6c src/HOL/Data_Structures/Trie_Map.thy --- a/src/HOL/Data_Structures/Trie_Map.thy Tue Sep 24 17:36:14 2019 +0200 +++ b/src/HOL/Data_Structures/Trie_Map.thy Wed Sep 25 17:22:57 2019 +0200 @@ -21,10 +21,10 @@ However, the development below works verbatim for any map implementation, eg \Tree_Map\, and not just \RBT_Map\, except for the termination lemma \lookup_size\.\ - +term size_tree lemma lookup_size[termination_simp]: fixes t :: "('a::linorder * 'a trie_map) rbt" - shows "lookup t a = Some b \ size b < Suc (size_tree (\ab. Suc (size (snd ab))) (\x. 0) t)" + shows "lookup t a = Some b \ size b < Suc (size_tree (\ab. Suc(Suc (size (snd(fst ab))))) t)" apply(induction t a rule: lookup.induct) apply(auto split: if_splits) done