--- a/src/HOL/Data_Structures/AA_Map.thy Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/AA_Map.thy Mon Jun 11 16:29:27 2018 +0200
@@ -9,21 +9,21 @@
begin
fun update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) aa_tree \<Rightarrow> ('a*'b) aa_tree" where
-"update x y Leaf = Node 1 Leaf (x,y) Leaf" |
-"update x y (Node lv t1 (a,b) 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 \<Rightarrow> split (skew (Node lv (update x y t1) (a,b) t2)) |
- GT \<Rightarrow> split (skew (Node lv t1 (a,b) (update x y t2))) |
- EQ \<Rightarrow> Node lv t1 (x,y) t2)"
+ LT \<Rightarrow> split (skew (Node (update x y t1) (a,b) lv t2)) |
+ GT \<Rightarrow> split (skew (Node t1 (a,b) lv (update x y t2))) |
+ EQ \<Rightarrow> Node t1 (x,y) lv t2)"
fun delete :: "'a::linorder \<Rightarrow> ('a*'b) aa_tree \<Rightarrow> ('a*'b) aa_tree" where
"delete _ Leaf = Leaf" |
-"delete x (Node lv l (a,b) r) =
+"delete x (Node l (a,b) lv r) =
(case cmp x a of
- LT \<Rightarrow> adjust (Node lv (delete x l) (a,b) r) |
- GT \<Rightarrow> adjust (Node lv l (a,b) (delete x r)) |
+ LT \<Rightarrow> adjust (Node (delete x l) (a,b) lv r) |
+ GT \<Rightarrow> adjust (Node l (a,b) lv (delete x r)) |
EQ \<Rightarrow> (if l = Leaf then r
- else let (l',ab') = split_max l in adjust (Node lv l' ab' r)))"
+ else let (l',ab') = split_max l in adjust (Node l' ab' lv r)))"
subsection "Invariance"
@@ -46,7 +46,7 @@
lemma lvl_update_sngl: "invar t \<Longrightarrow> sngl t \<Longrightarrow> lvl(update x y t) = lvl t"
proof (induction t rule: update.induct)
- case (2 x y lv t1 a b t2)
+ case (2 x y t1 a b lv t2)
consider (LT) "x < a" | (GT) "x > a" | (EQ) "x = a"
using less_linear by blast
thus ?case proof cases
@@ -64,7 +64,7 @@
qed simp
lemma lvl_update_incr_iff: "(lvl(update a b t) = lvl t + 1) \<longleftrightarrow>
- (\<exists>l x r. update a b t = Node (lvl t + 1) l x r \<and> lvl l = lvl r)"
+ (\<exists>l x r. update a b t = Node l x (lvl t + 1) r \<and> 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)
@@ -72,12 +72,12 @@
lemma invar_update: "invar t \<Longrightarrow> invar(update a b t)"
proof(induction t)
- case N: (Node n l xy r)
+ 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 n l xy r"
+ let ?t = "Node l xy n r"
have "a < x \<or> a = x \<or> x < a" by auto
moreover
have ?case if "a < x"
@@ -87,13 +87,13 @@
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 n t1 w 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 n (update a b l) xy r))"
+ have "update a b ?t = split(skew(Node (update a b l) xy n r))"
by(simp add: \<open>a<x\<close>)
- also have "skew(Node n (update a b l) xy r) = Node n t1 w (Node n t2 xy 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 \<dots>)"
proof (cases r)
@@ -101,7 +101,7 @@
hence "l = Leaf" using N.prems by(auto simp: lvl_0_iff)
thus ?thesis using Leaf ial by simp
next
- case [simp]: (Node m t3 y t4)
+ case [simp]: (Node t3 y m t4)
show ?thesis (*using N(3) iil l12 by(auto)*)
proof cases
assume "m = n" thus ?thesis using N(3) iil by(auto)
@@ -118,14 +118,14 @@
thus ?case
proof
assume 0: "n = lvl r"
- have "update a b ?t = split(skew(Node n l xy (update a b r)))"
+ have "update a b ?t = split(skew(Node l xy n (update a b r)))"
using \<open>a>x\<close> by(auto)
- also have "skew(Node n l xy (update a b r)) = Node n l xy (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 \<dots>)"
proof -
from lvl_update_sngl[OF ir sngl_if_invar[OF \<open>invar ?t\<close> 0], of a b]
- obtain t1 p t2 where iar: "update a b r = Node n t1 p 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)
@@ -157,12 +157,12 @@
theorem post_delete: "invar t \<Longrightarrow> post_del t (delete x t)"
proof (induction t)
- case (Node lv l ab r)
+ 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 lv l ab 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)
--- a/src/HOL/Data_Structures/AA_Set.thy Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/AA_Set.thy Mon Jun 11 16:29:27 2018 +0200
@@ -14,60 +14,60 @@
fun lvl :: "'a aa_tree \<Rightarrow> nat" where
"lvl Leaf = 0" |
-"lvl (Node lv _ _ _) = lv"
+"lvl (Node _ _ lv _) = lv"
fun invar :: "'a aa_tree \<Rightarrow> bool" where
"invar Leaf = True" |
-"invar (Node h l a r) =
+"invar (Node l a h r) =
(invar l \<and> invar r \<and>
- h = lvl l + 1 \<and> (h = lvl r + 1 \<or> (\<exists>lr b rr. r = Node h lr b rr \<and> h = lvl rr + 1)))"
+ h = lvl l + 1 \<and> (h = lvl r + 1 \<or> (\<exists>lr b rr. r = Node lr b h rr \<and> h = lvl rr + 1)))"
fun skew :: "'a aa_tree \<Rightarrow> 'a aa_tree" where
-"skew (Node lva (Node lvb t1 b t2) a t3) =
- (if lva = lvb then Node lva t1 b (Node lva t2 a t3) else Node lva (Node lvb t1 b t2) a 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 \<Rightarrow> 'a aa_tree" where
-"split (Node lva t1 a (Node lvb t2 b (Node lvc t3 c t4))) =
+"split (Node t1 a lva (Node t2 b lvb (Node t3 c lvc t4))) =
(if lva = lvb \<and> lvb = lvc \<comment> \<open>\<open>lva = lvc\<close> suffices\<close>
- then Node (lva+1) (Node lva t1 a t2) b (Node lva t3 c t4)
- else Node lva t1 a (Node lvb t2 b (Node lvc t3 c 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 \<Rightarrow> 'a aa_tree \<Rightarrow> 'a aa_tree" where
-"insert x Leaf = Node 1 Leaf x Leaf" |
-"insert x (Node lv t1 a t2) =
+"insert x Leaf = Node Leaf x 1 Leaf" |
+"insert x (Node t1 a lv t2) =
(case cmp x a of
- LT \<Rightarrow> split (skew (Node lv (insert x t1) a t2)) |
- GT \<Rightarrow> split (skew (Node lv t1 a (insert x t2))) |
- EQ \<Rightarrow> Node lv t1 x t2)"
+ LT \<Rightarrow> split (skew (Node (insert x t1) a lv t2)) |
+ GT \<Rightarrow> split (skew (Node t1 a lv (insert x t2))) |
+ EQ \<Rightarrow> Node t1 x lv t2)"
fun sngl :: "'a aa_tree \<Rightarrow> bool" where
"sngl Leaf = False" |
"sngl (Node _ _ _ Leaf) = True" |
-"sngl (Node lva _ _ (Node lvb _ _ _)) = (lva > lvb)"
+"sngl (Node _ _ lva (Node _ _ lvb _)) = (lva > lvb)"
definition adjust :: "'a aa_tree \<Rightarrow> 'a aa_tree" where
"adjust t =
(case t of
- Node lv l x r \<Rightarrow>
+ Node l x lv r \<Rightarrow>
(if lvl l >= lv-1 \<and> lvl r >= lv-1 then t else
- if lvl r < lv-1 \<and> sngl l then skew (Node (lv-1) l x r) else
+ if lvl r < lv-1 \<and> sngl l then skew (Node l x (lv-1) r) else
if lvl r < lv-1
then case l of
- Node lva t1 a (Node lvb t2 b t3)
- \<Rightarrow> Node (lvb+1) (Node lva t1 a t2) b (Node (lv-1) t3 x r)
+ Node t1 a lva (Node t2 b lvb t3)
+ \<Rightarrow> Node (Node t1 a lva t2) b (lvb+1) (Node t3 x (lv-1) r)
else
- if lvl r < lv then split (Node (lv-1) l x r)
+ if lvl r < lv then split (Node l x (lv-1) r)
else
case r of
- Node lvb t1 b t4 \<Rightarrow>
+ Node t1 b lvb t4 \<Rightarrow>
(case t1 of
- Node lva t2 a t3
- \<Rightarrow> Node (lva+1) (Node (lv-1) l x t2) a
- (split (Node (if sngl t1 then lva else lva+1) t3 b t4)))))"
+ Node t2 a lva t3
+ \<Rightarrow> Node (Node l x (lv-1) t2) a (lva+1)
+ (split (Node t3 b (if sngl t1 then lva else lva+1) t4)))))"
text\<open>In the paper, the last case of @{const adjust} is expressed with the help of an
incorrect auxiliary function \texttt{nlvl}.
@@ -78,20 +78,20 @@
is not restored.\<close>
fun split_max :: "'a aa_tree \<Rightarrow> 'a aa_tree * 'a" where
-"split_max (Node lv l a Leaf) = (l,a)" |
-"split_max (Node lv l a r) = (let (r',b) = split_max r in (adjust(Node lv l a 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 \<Rightarrow> 'a aa_tree \<Rightarrow> 'a aa_tree" where
"delete _ Leaf = Leaf" |
-"delete x (Node lv l a r) =
+"delete x (Node l a lv r) =
(case cmp x a of
- LT \<Rightarrow> adjust (Node lv (delete x l) a r) |
- GT \<Rightarrow> adjust (Node lv l a (delete x r)) |
+ LT \<Rightarrow> adjust (Node (delete x l) a lv r) |
+ GT \<Rightarrow> adjust (Node l a lv (delete x r)) |
EQ \<Rightarrow> (if l = Leaf then r
- else let (l',b) = split_max l in adjust (Node lv l' b r)))"
+ else let (l',b) = split_max l in adjust (Node l' b lv r)))"
fun pre_adjust where
-"pre_adjust (Node lv l a r) = (invar l \<and> invar r \<and>
+"pre_adjust (Node l a lv r) = (invar l \<and> invar r \<and>
((lv = lvl l + 1 \<and> (lv = lvl r + 1 \<or> lv = lvl r + 2 \<or> lv = lvl r \<and> sngl r)) \<or>
(lv = lvl l + 2 \<and> (lv = lvl r + 1 \<or> lv = lvl r \<and> sngl r))))"
@@ -100,23 +100,23 @@
subsection "Auxiliary Proofs"
lemma split_case: "split t = (case t of
- Node lvx a x (Node lvy b y (Node lvz c z d)) \<Rightarrow>
+ Node t1 x lvx (Node t2 y lvy (Node t3 z lvz t4)) \<Rightarrow>
(if lvx = lvy \<and> lvy = lvz
- then Node (lvx+1) (Node lvx a x b) y (Node lvx c z d)
+ then Node (Node t1 x lvx t2) y (lvx+1) (Node t3 z lvx t4)
else t)
| t \<Rightarrow> t)"
by(auto split: tree.split)
lemma skew_case: "skew t = (case t of
- Node lvx (Node lvy a y b) x c \<Rightarrow>
- (if lvx = lvy then Node lvx a y (Node lvx b x c) else t)
+ Node (Node t1 y lvy t2) x lvx t3 \<Rightarrow>
+ (if lvx = lvy then Node t1 y lvx (Node t2 x lvx t3) else t)
| t \<Rightarrow> t)"
by(auto split: tree.split)
lemma lvl_0_iff: "invar t \<Longrightarrow> lvl t = 0 \<longleftrightarrow> t = Leaf"
by(cases t) auto
-lemma lvl_Suc_iff: "lvl t = Suc n \<longleftrightarrow> (\<exists> l a r. t = Node (Suc n) l a r)"
+lemma lvl_Suc_iff: "lvl t = Suc n \<longleftrightarrow> (\<exists> l a r. t = Node l a (Suc n) r)"
by(cases t) auto
lemma lvl_skew: "lvl (skew t) = lvl t"
@@ -125,16 +125,16 @@
lemma lvl_split: "lvl (split t) = lvl t \<or> lvl (split t) = lvl t + 1 \<and> sngl (split t)"
by(cases t rule: split.cases) auto
-lemma invar_2Nodes:"invar (Node lv l x (Node rlv rl rx rr)) =
- (invar l \<and> invar \<langle>rlv, rl, rx, rr\<rangle> \<and> lv = Suc (lvl l) \<and>
+lemma invar_2Nodes:"invar (Node l x lv (Node rl rx rlv rr)) =
+ (invar l \<and> invar \<langle>rl, rx, rlv, rr\<rangle> \<and> lv = Suc (lvl l) \<and>
(lv = Suc rlv \<or> rlv = lv \<and> lv = Suc (lvl rr)))"
by simp
lemma invar_NodeLeaf[simp]:
- "invar (Node lv l x Leaf) = (invar l \<and> lv = Suc (lvl l) \<and> lv = Suc 0)"
+ "invar (Node l x lv Leaf) = (invar l \<and> lv = Suc (lvl l) \<and> lv = Suc 0)"
by simp
-lemma sngl_if_invar: "invar (Node n l a r) \<Longrightarrow> n = lvl r \<Longrightarrow> sngl r"
+lemma sngl_if_invar: "invar (Node l a n r) \<Longrightarrow> n = lvl r \<Longrightarrow> sngl r"
by(cases r rule: sngl.cases) clarsimp+
@@ -156,7 +156,7 @@
lemma lvl_insert_sngl: "invar t \<Longrightarrow> sngl t \<Longrightarrow> lvl(insert x t) = lvl t"
proof (induction t rule: insert.induct)
- case (2 x lv t1 a t2)
+ case (2 x t1 a lv t2)
consider (LT) "x < a" | (GT) "x > a" | (EQ) "x = a"
using less_linear by blast
thus ?case proof cases
@@ -180,20 +180,20 @@
by(cases t rule: split.cases) clarsimp+
lemma invar_NodeL:
- "\<lbrakk> invar(Node n l x r); invar l'; lvl l' = lvl l \<rbrakk> \<Longrightarrow> invar(Node n l' x r)"
+ "\<lbrakk> invar(Node l x n r); invar l'; lvl l' = lvl l \<rbrakk> \<Longrightarrow> invar(Node l' x n r)"
by(auto)
lemma invar_NodeR:
- "\<lbrakk> invar(Node n l x r); n = lvl r + 1; invar r'; lvl r' = lvl r \<rbrakk> \<Longrightarrow> invar(Node n l x r')"
+ "\<lbrakk> invar(Node l x n r); n = lvl r + 1; invar r'; lvl r' = lvl r \<rbrakk> \<Longrightarrow> invar(Node l x n r')"
by(auto)
lemma invar_NodeR2:
- "\<lbrakk> invar(Node n l x r); sngl r'; n = lvl r + 1; invar r'; lvl r' = n \<rbrakk> \<Longrightarrow> invar(Node n l x r')"
+ "\<lbrakk> invar(Node l x n r); sngl r'; n = lvl r + 1; invar r'; lvl r' = n \<rbrakk> \<Longrightarrow> 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) \<longleftrightarrow>
- (\<exists>l x r. insert a t = Node (lvl t + 1) l x r \<and> lvl l = lvl r)"
+ (\<exists>l x r. insert a t = Node l x (lvl t + 1) r \<and> 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)
@@ -201,11 +201,11 @@
lemma invar_insert: "invar t \<Longrightarrow> invar(insert a t)"
proof(induction t)
- case N: (Node n l x r)
+ 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 n l x r"
+ let ?t = "Node l x n r"
have "a < x \<or> a = x \<or> x < a" by auto
moreover
have ?case if "a < x"
@@ -215,13 +215,13 @@
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 n t1 w 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 n (insert a l) x r))"
+ have "insert a ?t = split(skew(Node (insert a l) x n r))"
by(simp add: \<open>a<x\<close>)
- also have "skew(Node n (insert a l) x r) = Node n t1 w (Node n t2 x 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 \<dots>)"
proof (cases r)
@@ -229,7 +229,7 @@
hence "l = Leaf" using N.prems by(auto simp: lvl_0_iff)
thus ?thesis using Leaf ial by simp
next
- case [simp]: (Node m t3 y t4)
+ case [simp]: (Node t3 y m t4)
show ?thesis (*using N(3) iil l12 by(auto)*)
proof cases
assume "m = n" thus ?thesis using N(3) iil by(auto)
@@ -246,14 +246,14 @@
thus ?case
proof
assume 0: "n = lvl r"
- have "insert a ?t = split(skew(Node n l x (insert a r)))"
+ have "insert a ?t = split(skew(Node l x n (insert a r)))"
using \<open>a>x\<close> by(auto)
- also have "skew(Node n l x (insert a r)) = Node n l x (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 \<dots>)"
proof -
from lvl_insert_sngl[OF ir sngl_if_invar[OF \<open>invar ?t\<close> 0], of a]
- obtain t1 y t2 where iar: "insert a r = Node n t1 y 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)
@@ -282,21 +282,21 @@
subsubsection "Proofs for delete"
-lemma invarL: "ASSUMPTION(invar \<langle>lv, l, a, r\<rangle>) \<Longrightarrow> invar l"
+lemma invarL: "ASSUMPTION(invar \<langle>l, a, lv, r\<rangle>) \<Longrightarrow> invar l"
by(simp add: ASSUMPTION_def)
lemma invarR: "ASSUMPTION(invar \<langle>lv, l, a, r\<rangle>) \<Longrightarrow> invar r"
by(simp add: ASSUMPTION_def)
lemma sngl_NodeI:
- "sngl (Node lv l a r) \<Longrightarrow> sngl (Node lv l' a' r)"
+ "sngl (Node l a lv r) \<Longrightarrow> sngl (Node l' a' lv r)"
by(cases r) (simp_all)
declare invarL[simp] invarR[simp]
lemma pre_cases:
-assumes "pre_adjust (Node lv l x r)"
+assumes "pre_adjust (Node l x lv r)"
obtains
(tSngl) "invar l \<and> invar r \<and>
lv = Suc (lvl r) \<and> lvl l = lvl r" |
@@ -314,38 +314,38 @@
declare invar.simps(2)[simp del] invar_2Nodes[simp add]
lemma invar_adjust:
- assumes pre: "pre_adjust (Node lv l a r)"
- shows "invar(adjust (Node lv l a 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 llv ll la 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 rlv rl ra 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 rrlv rrr rra 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)
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 lv l a r)"
- shows "lv = lvl (adjust(Node lv l a r)) \<or> lv = lvl (adjust(Node lv l a r)) + 1"
+ assumes "pre_adjust (Node l a lv r)"
+ shows "lv = lvl (adjust(Node l a lv r)) \<or> 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 "\<langle>lvl r, l, a, r\<rangle>"] by (auto simp: adjust_def)
+ using lvl_split[of "\<langle>l, a, lvl r, r\<rangle>"] 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 lv l a r)"
- "sngl \<langle>lv, l, a, r\<rangle>" "lv = lvl (adjust \<langle>lv, l, a, r\<rangle>)"
- shows "sngl (adjust \<langle>lv, l, a, r\<rangle>)"
+lemma sngl_adjust: assumes "pre_adjust (Node l a lv r)"
+ "sngl \<langle>l, a, lv, r\<rangle>" "lv = lvl (adjust \<langle>l, a, lv, r\<rangle>)"
+ shows "sngl (adjust \<langle>l, a, lv, r\<rangle>)"
using assms proof (cases rule: pre_cases)
case rDown
thus ?thesis using assms(2,3) unfolding adjust_def
@@ -363,13 +363,13 @@
(auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims)
lemma pre_adj_if_postL:
- "invar\<langle>lv, l, a, r\<rangle> \<Longrightarrow> post_del l l' \<Longrightarrow> pre_adjust \<langle>lv, l', b, r\<rangle>"
+ "invar\<langle>l, a, lv, r\<rangle> \<Longrightarrow> post_del l l' \<Longrightarrow> pre_adjust \<langle>l', b, lv, r\<rangle>"
by(cases "sngl r")
(auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims)
lemma post_del_adjL:
- "\<lbrakk> invar\<langle>lv, l, a, r\<rangle>; pre_adjust \<langle>lv, l', b, r\<rangle> \<rbrakk>
- \<Longrightarrow> post_del \<langle>lv, l, a, r\<rangle> (adjust \<langle>lv, l', b, r\<rangle>)"
+ "\<lbrakk> invar\<langle>l, a, lv, r\<rangle>; pre_adjust \<langle>l', b, lv, r\<rangle> \<rbrakk>
+ \<Longrightarrow> post_del \<langle>l, a, lv, r\<rangle> (adjust \<langle>l', b, lv, r\<rangle>)"
unfolding post_del_def
by (metis invar_adjust lvl_adjust sngl_NodeI sngl_adjust lvl.simps(2))
@@ -412,10 +412,10 @@
theorem post_delete: "invar t \<Longrightarrow> post_del t (delete x t)"
proof (induction t)
- case (Node lv l a r)
+ case (Node l a lv r)
let ?l' = "delete x l" and ?r' = "delete x r"
- let ?t = "Node lv l a 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)
--- a/src/HOL/Data_Structures/AVL_Map.thy Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/AVL_Map.thy Mon Jun 11 16:29:27 2018 +0200
@@ -9,16 +9,16 @@
begin
fun update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) avl_tree \<Rightarrow> ('a*'b) avl_tree" where
-"update x y Leaf = Node 1 Leaf (x,y) Leaf" |
-"update x y (Node h l (a,b) r) = (case cmp x a of
- EQ \<Rightarrow> Node h l (x,y) 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 \<Rightarrow> Node l (x,y) h r |
LT \<Rightarrow> balL (update x y l) (a,b) r |
GT \<Rightarrow> balR l (a,b) (update x y r))"
fun delete :: "'a::linorder \<Rightarrow> ('a*'b) avl_tree \<Rightarrow> ('a*'b) avl_tree" where
"delete _ Leaf = Leaf" |
-"delete x (Node h l (a,b) r) = (case cmp x a of
- EQ \<Rightarrow> del_root (Node h l (a,b) r) |
+"delete x (Node l (a,b) h r) = (case cmp x a of
+ EQ \<Rightarrow> del_root (Node l (a,b) h r) |
LT \<Rightarrow> balR (delete x l) (a,b) r |
GT \<Rightarrow> balL l (a,b) (delete x r))"
--- a/src/HOL/Data_Structures/AVL_Set.thy Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/AVL_Set.thy Mon Jun 11 16:29:27 2018 +0200
@@ -18,25 +18,25 @@
fun avl :: "'a avl_tree \<Rightarrow> bool" where
"avl Leaf = True" |
-"avl (Node h l a r) =
+"avl (Node l a h r) =
((height l = height r \<or> height l = height r + 1 \<or> height r = height l + 1) \<and>
h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)"
fun ht :: "'a avl_tree \<Rightarrow> nat" where
"ht Leaf = 0" |
-"ht (Node h l a r) = h"
+"ht (Node l a h r) = h"
definition node :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
-"node l a r = Node (max (ht l) (ht r) + 1) l a r"
+"node l a r = Node l a (max (ht l) (ht r) + 1) r"
definition balL :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
"balL l a r =
(if ht l = ht r + 2 then
case l of
- Node _ bl b br \<Rightarrow>
+ Node bl b _ br \<Rightarrow>
if ht bl < ht br then
case br of
- Node _ cl c cr \<Rightarrow> node (node bl b cl) c (node cr a r)
+ Node cl c _ cr \<Rightarrow> node (node bl b cl) c (node cr a r)
else node bl b (node br a r)
else node l a r)"
@@ -44,38 +44,38 @@
"balR l a r =
(if ht r = ht l + 2 then
case r of
- Node _ bl b br \<Rightarrow>
+ Node bl b _ br \<Rightarrow>
if ht bl > ht br then
case bl of
- Node _ cl c cr \<Rightarrow> node (node l a cl) c (node cr b br)
+ Node cl c _ cr \<Rightarrow> 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 \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
-"insert x Leaf = Node 1 Leaf x Leaf" |
-"insert x (Node h l a r) = (case cmp x a of
- EQ \<Rightarrow> Node h l a r |
+"insert x Leaf = Node Leaf x 1 Leaf" |
+"insert x (Node l a h r) = (case cmp x a of
+ EQ \<Rightarrow> Node l a h r |
LT \<Rightarrow> balL (insert x l) a r |
GT \<Rightarrow> balR l a (insert x r))"
fun split_max :: "'a avl_tree \<Rightarrow> '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 \<Rightarrow> 'a avl_tree" where
-"del_root (Node h Leaf a r) = r" |
-"del_root (Node h l a Leaf) = l" |
-"del_root (Node h l a 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]
fun delete :: "'a::linorder \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where
"delete _ Leaf = Leaf" |
-"delete x (Node h l a r) =
+"delete x (Node l a h r) =
(case cmp x a of
- EQ \<Rightarrow> del_root (Node h l a r) |
+ EQ \<Rightarrow> del_root (Node l a h r) |
LT \<Rightarrow> balR (delete x l) a r |
GT \<Rightarrow> balL l a (delete x r))"
@@ -110,8 +110,8 @@
(auto simp: inorder_balL split: if_splits prod.splits tree.split)
lemma inorder_del_root:
- "inorder (del_root (Node h l a r)) = inorder l @ inorder r"
-by(cases "Node h l a r" rule: del_root.cases)
+ "inorder (del_root (Node l a h r)) = inorder l @ inorder r"
+by(cases "Node l a h r" rule: del_root.cases)
(auto simp: inorder_balL inorder_balR inorder_split_maxD split: if_splits prod.splits)
theorem inorder_delete:
@@ -188,7 +188,7 @@
case Leaf
with assms show ?thesis by (simp add: node_def balL_def)
next
- case (Node ln ll lr lh)
+ case Node
with assms show ?thesis
proof(cases "height l = height r + 2")
case True
@@ -208,7 +208,7 @@
case Leaf
with assms show ?thesis by (simp add: node_def balR_def)
next
- case (Node rn rl rr rh)
+ case Node
with assms show ?thesis
proof(cases "height r = height l + 2")
case True
@@ -230,7 +230,7 @@
"(height (insert x t) = height t \<or> height (insert x t) = height t + 1)"
using assms
proof (induction t)
- case (Node h l a r)
+ case (Node l a h r)
case 1
with Node show ?case
proof(cases "x = a")
@@ -307,14 +307,14 @@
height x = height(fst (split_max x)) + 1"
using assms
proof (induct x rule: split_max_induct)
- case (Node h l a r)
+ case (Node l a h r)
case 1
thus ?case using Node
by (auto simp: height_balL height_balL2 avl_balL
linorder_class.max.absorb1 linorder_class.max.absorb2
split:prod.split)
next
- case (Node h l a r)
+ case (Node l a h r)
case 2
let ?r' = "fst (split_max r)"
from \<open>avl x\<close> Node 2 have "avl l" and "avl r" by simp_all
@@ -327,9 +327,9 @@
shows "avl(del_root t)"
using assms
proof (cases t rule:del_root_cases)
- case (Node_Node h lh ll ln lr n rh rl rn rr)
- let ?l = "Node lh ll ln lr"
- let ?r = "Node rh rl rn rr"
+ 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' = "fst (split_max ?l)"
from \<open>avl t\<close> and Node_Node have "avl ?r" by simp
from \<open>avl t\<close> and Node_Node have "avl ?l" by simp
@@ -347,9 +347,9 @@
shows "height t = height(del_root t) \<or> height t = height(del_root t) + 1"
using assms
proof (cases t rule: del_root_cases)
- case (Node_Node h lh ll ln lr n rh rl rn rr)
- let ?l = "Node lh ll ln lr"
- let ?r = "Node rh rl rn rr"
+ 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' = "fst (split_max ?l)"
let ?t' = "balR ?l' (snd(split_max ?l)) ?r"
from \<open>avl t\<close> and Node_Node have "avl ?r" by simp
@@ -382,7 +382,7 @@
shows "avl(delete x t)" and "height t = (height (delete x t)) \<or> height t = height (delete x t) + 1"
using assms
proof (induct t)
- case (Node h l n r)
+ case (Node l n h r)
case 1
with Node show ?case
proof(cases "x = n")
@@ -403,8 +403,8 @@
with Node show ?case
proof(cases "x = n")
case True
- with 1 have "height (Node h l n r) = height(del_root (Node h l n r))
- \<or> height (Node h l n r) = height(del_root (Node h l n r)) + 1"
+ with 1 have "height (Node l n h r) = height(del_root (Node l n h r))
+ \<or> 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
@@ -459,7 +459,7 @@
lemma height_invers:
"(height t = 0) = (t = Leaf)"
- "avl t \<Longrightarrow> (height t = Suc h) = (\<exists> l a r . t = Node (Suc h) l a r)"
+ "avl t \<Longrightarrow> (height t = Suc h) = (\<exists> l a r . t = Node l a (Suc h) r)"
by (induction t) auto
text \<open>Any AVL tree of height \<open>h\<close> has at least \<open>fib (h+2)\<close> leaves:\<close>
@@ -472,7 +472,7 @@
next
case (3 h)
from "3.prems" obtain l a r where
- [simp]: "t = Node (Suc(Suc h)) l a r" "avl l" "avl r"
+ [simp]: "t = Node l a (Suc(Suc h)) r" "avl l" "avl r"
and C: "
height r = Suc h \<and> height l = Suc h
\<or> height r = Suc h \<and> height l = h
--- a/src/HOL/Data_Structures/Isin2.thy Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/Isin2.thy Mon Jun 11 16:29:27 2018 +0200
@@ -11,7 +11,7 @@
fun isin :: "('a::linorder,'b) tree \<Rightarrow> 'a \<Rightarrow> bool" where
"isin Leaf x = False" |
-"isin (Node _ l a r) x =
+"isin (Node l a _ r) x =
(case cmp x a of
LT \<Rightarrow> isin l x |
EQ \<Rightarrow> True |
--- a/src/HOL/Data_Structures/Leftist_Heap.thy Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/Leftist_Heap.thy Mon Jun 11 16:29:27 2018 +0200
@@ -12,7 +12,7 @@
fun mset_tree :: "('a,'b) tree \<Rightarrow> '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"
@@ -22,27 +22,27 @@
fun rk :: "'a lheap \<Rightarrow> nat" where
"rk Leaf = 0" |
-"rk (Node n _ _ _) = n"
+"rk (Node _ _ n _) = n"
text\<open>The invariants:\<close>
fun (in linorder) heap :: "('a,'b) tree \<Rightarrow> bool" where
"heap Leaf = True" |
-"heap (Node _ l m r) =
+"heap (Node l m _ r) =
(heap l \<and> heap r \<and> (\<forall>x \<in> set_mset(mset_tree l + mset_tree r). m \<le> x))"
fun ltree :: "'a lheap \<Rightarrow> bool" where
"ltree Leaf = True" |
-"ltree (Node n l a r) =
+"ltree (Node l a n r) =
(n = rank r + 1 \<and> rank l \<ge> rank r \<and> ltree l & ltree r)"
definition node :: "'a lheap \<Rightarrow> 'a \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap" where
"node l a r =
(let rl = rk l; rr = rk r
- in if rl \<ge> rr then Node (rr+1) l a r else Node (rl+1) r a l)"
+ in if rl \<ge> rr then Node l a (rr+1) r else Node r a (rl+1) l)"
fun get_min :: "'a lheap \<Rightarrow> 'a" where
-"get_min(Node n l a r) = a"
+"get_min(Node l a n r) = a"
text \<open>For function \<open>merge\<close>:\<close>
unbundle pattern_aliases
@@ -51,25 +51,25 @@
fun merge :: "'a::ord lheap \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap" where
"merge Leaf t2 = t2" |
"merge t1 Leaf = t1" |
-"merge (Node n1 l1 a1 r1 =: t1) (Node n2 l2 a2 r2 =: t2) =
+"merge (Node l1 a1 n1 r1 =: t1) (Node l2 a2 n2 r2 =: t2) =
(if a1 \<le> a2 then node l1 a1 (merge r1 t2)
else node l2 a2 (merge r2 t1))"
lemma merge_code: "merge t1 t2 = (case (t1,t2) of
(Leaf, _) \<Rightarrow> t2 |
(_, Leaf) \<Rightarrow> t1 |
- (Node n1 l1 a1 r1, Node n2 l2 a2 r2) \<Rightarrow>
+ (Node l1 a1 n1 r1, Node l2 a2 n2 r2) \<Rightarrow>
if a1 \<le> a2 then node l1 a1 (merge r1 t2) else node l2 a2 (merge r2 t1))"
by(induction t1 t2 rule: merge.induct) (simp_all split: tree.split)
hide_const (open) insert
definition insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> 'a lheap" where
-"insert x t = merge (Node 1 Leaf x Leaf) t"
+"insert x t = merge (Node Leaf x 1 Leaf) t"
fun del_min :: "'a::ord lheap \<Rightarrow> 'a lheap" where
"del_min Leaf = Leaf" |
-"del_min (Node n l x r) = merge l r"
+"del_min (Node l x n r) = merge l r"
subsection "Lemmas"
@@ -104,7 +104,7 @@
lemma ltree_merge: "\<lbrakk> ltree l; ltree r \<rbrakk> \<Longrightarrow> ltree (merge l r)"
proof(induction l r rule: merge.induct)
- case (3 n1 l1 a1 r1 n2 l2 a2 r2)
+ case (3 l1 a1 n1 r1 l2 a2 n2 r2)
show ?case (is "ltree(merge ?t1 ?t2)")
proof cases
assume "a1 \<le> a2"
@@ -173,14 +173,14 @@
proof(induction t)
case Leaf show ?case by simp
next
- case (Node n l a r)
+ case (Node l a n r)
hence "rank r \<le> rank l" by simp
hence *: "(2::nat) ^ rank r \<le> 2 ^ rank l" by simp
- have "(2::nat) ^ rank \<langle>n, l, a, r\<rangle> = 2 ^ rank r + 2 ^ rank r"
+ have "(2::nat) ^ rank \<langle>l, a, n, r\<rangle> = 2 ^ rank r + 2 ^ rank r"
by(simp add: mult_2)
also have "\<dots> \<le> size1 l + size1 r"
using Node * by (simp del: power_increasing_iff)
- also have "\<dots> = size1 \<langle>n, l, a, r\<rangle>" by simp
+ also have "\<dots> = size1 \<langle>l, a, n, r\<rangle>" by simp
finally show ?case .
qed
@@ -189,16 +189,16 @@
fun t_merge :: "'a::ord lheap \<Rightarrow> 'a lheap \<Rightarrow> nat" where
"t_merge Leaf t2 = 1" |
"t_merge t2 Leaf = 1" |
-"t_merge (Node n1 l1 a1 r1 =: t1) (Node n2 l2 a2 r2 =: t2) =
+"t_merge (Node l1 a1 n1 r1 =: t1) (Node l2 a2 n2 r2 =: t2) =
(if a1 \<le> a2 then 1 + t_merge r1 t2
else 1 + t_merge r2 t1)"
definition t_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> nat" where
-"t_insert x t = t_merge (Node 1 Leaf x Leaf) t"
+"t_insert x t = t_merge (Node Leaf x 1 Leaf) t"
fun t_del_min :: "'a::ord lheap \<Rightarrow> nat" where
"t_del_min Leaf = 1" |
-"t_del_min (Node n l a r) = t_merge l r"
+"t_del_min (Node l a n r) = t_merge l r"
lemma t_merge_rank: "t_merge l r \<le> rank l + rank r + 1"
proof(induction l r rule: merge.induct)
@@ -213,7 +213,7 @@
by linarith
corollary t_insert_log: "ltree t \<Longrightarrow> t_insert x t \<le> log 2 (size1 t) + 2"
-using t_merge_log[of "Node 1 Leaf x Leaf" t]
+using t_merge_log[of "Node Leaf x 1 Leaf" t]
by(simp add: t_insert_def split: tree.split)
(* FIXME mv ? *)
@@ -234,7 +234,7 @@
proof(cases t)
case Leaf thus ?thesis using assms by simp
next
- case [simp]: (Node _ t1 _ t2)
+ case [simp]: (Node t1 _ _ t2)
have "t_del_min t = t_merge t1 t2" by simp
also have "\<dots> \<le> log 2 (size1 t1) + log 2 (size1 t2) + 1"
using \<open>ltree t\<close> by (auto simp: t_merge_log simp del: t_merge.simps)
--- a/src/HOL/Data_Structures/Lookup2.thy Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/Lookup2.thy Mon Jun 11 16:29:27 2018 +0200
@@ -11,7 +11,7 @@
fun lookup :: "('a::linorder * 'b, 'c) tree \<Rightarrow> 'a \<Rightarrow> '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 \<Rightarrow> lookup l x | GT \<Rightarrow> lookup r x | EQ \<Rightarrow> Some b)"
lemma lookup_map_of:
--- a/src/HOL/Data_Structures/RBT.thy Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/RBT.thy Mon Jun 11 16:29:27 2018 +0200
@@ -10,8 +10,8 @@
type_synonym 'a rbt = "('a,color)tree"
-abbreviation R where "R l a r \<equiv> Node Red l a r"
-abbreviation B where "B l a r \<equiv> Node Black l a r"
+abbreviation R where "R l a r \<equiv> Node l a Red r"
+abbreviation B where "B l a r \<equiv> Node l a Black r"
fun baliL :: "'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
"baliL (R (R t1 a1 t2) a2 t3) a3 t4 = R (B t1 a1 t2) a2 (B t3 a3 t4)" |
@@ -25,7 +25,7 @@
fun paint :: "color \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
"paint c Leaf = Leaf" |
-"paint c (Node _ l a r) = Node c l a r"
+"paint c (Node l a _ r) = Node l a c r"
fun baldL :: "'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
"baldL (R t1 x t2) y t3 = R (B t1 x t2) y t3" |
--- a/src/HOL/Data_Structures/RBT_Map.thy Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/RBT_Map.thy Mon Jun 11 16:29:27 2018 +0200
@@ -27,7 +27,7 @@
and delR :: "'a::linorder \<Rightarrow> ('a*'b)rbt \<Rightarrow> 'a*'b \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
where
"del x Leaf = Leaf" |
-"del x (Node c t1 (a,b) t2) = (case cmp x a of
+"del x (Node t1 (a,b) c t2) = (case cmp x a of
LT \<Rightarrow> delL x t1 (a,b) t2 |
GT \<Rightarrow> delR x t1 (a,b) t2 |
EQ \<Rightarrow> combine t1 t2)" |
--- a/src/HOL/Data_Structures/RBT_Set.thy Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/RBT_Set.thy Mon Jun 11 16:29:27 2018 +0200
@@ -28,11 +28,11 @@
fun color :: "'a rbt \<Rightarrow> color" where
"color Leaf = Black" |
-"color (Node c _ _ _) = c"
+"color (Node _ _ c _) = c"
fun del :: "'a::linorder \<Rightarrow> 'a rbt \<Rightarrow> '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 \<Rightarrow> if l \<noteq> Leaf \<and> color l = Black
then baldL (del x l) a r else R (del x l) a r |
@@ -97,20 +97,20 @@
fun bheight :: "'a rbt \<Rightarrow> nat" where
"bheight Leaf = 0" |
-"bheight (Node c l x 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 \<Rightarrow> bool" where
"invc Leaf = True" |
-"invc (Node c l a r) =
+"invc (Node l a c r) =
(invc l \<and> invc r \<and> (c = Red \<longrightarrow> color l = Black \<and> color r = Black))"
fun invc2 :: "'a rbt \<Rightarrow> bool" \<comment> \<open>Weaker version\<close> where
"invc2 Leaf = True" |
-"invc2 (Node c l a r) = (invc l \<and> invc r)"
+"invc2 (Node l a c r) = (invc l \<and> invc r)"
fun invh :: "'a rbt \<Rightarrow> bool" where
"invh Leaf = True" |
-"invh (Node c l x r) = (invh l \<and> invh r \<and> bheight l = bheight r)"
+"invh (Node l x c r) = (invh l \<and> invh r \<and> bheight l = bheight r)"
lemma invc2I: "invc t \<Longrightarrow> invc2 t"
by (cases t) simp+
@@ -232,7 +232,7 @@
(color t = Red \<and> bheight (del x t) = bheight t \<and> invc (del x t) \<or>
color t = Black \<and> bheight (del x t) = bheight t - 1 \<and> invc2 (del x t))"
proof (induct x t rule: del.induct)
-case (2 x c _ y)
+case (2 x _ y c)
have "x = y \<or> x < y \<or> x > y" by auto
thus ?case proof (elim disjE)
assume "x = y"
--- a/src/HOL/Data_Structures/Set2_Join.thy Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/Set2_Join.thy Mon Jun 11 16:29:27 2018 +0200
@@ -28,7 +28,7 @@
\<Longrightarrow> bst (join l a r)"
assumes inv_Leaf: "inv \<langle>\<rangle>"
assumes inv_join: "\<lbrakk> inv l; inv r \<rbrakk> \<Longrightarrow> inv (join l k r)"
-assumes inv_Node: "\<lbrakk> inv (Node h l x r) \<rbrakk> \<Longrightarrow> inv l \<and> inv r"
+assumes inv_Node: "\<lbrakk> inv (Node l x h r) \<rbrakk> \<Longrightarrow> inv l \<and> inv r"
begin
declare set_join [simp]
@@ -36,7 +36,7 @@
subsection "\<open>split_min\<close>"
fun split_min :: "('a,'b) tree \<Rightarrow> 'a \<times> ('a,'b) tree" where
-"split_min (Node _ l x r) =
+"split_min (Node l x _ r) =
(if l = Leaf then (x,r) else let (m,l') = split_min l in (m, join l' x r))"
lemma split_min_set:
@@ -84,7 +84,7 @@
fun split :: "('a,'b)tree \<Rightarrow> 'a \<Rightarrow> ('a,'b)tree \<times> bool \<times> ('a,'b)tree" where
"split Leaf k = (Leaf, False, Leaf)" |
-"split (Node _ l a r) k =
+"split (Node l a _ r) k =
(if k < a then let (l1,b,l2) = split l k in (l1, b, join l2 a r) else
if a < k then let (r1,b,r2) = split r k in (join l a r1, b, r2)
else (l, True, r))"
@@ -145,7 +145,7 @@
"union t1 t2 =
(if t1 = Leaf then t2 else
if t2 = Leaf then t1 else
- case t1 of Node _ l1 k r1 \<Rightarrow>
+ case t1 of Node l1 k _ r1 \<Rightarrow>
let (l2,_ ,r2) = split t2 k;
l' = union l1 l2; r' = union r1 r2
in join l' k r')"
@@ -181,7 +181,7 @@
"inter t1 t2 =
(if t1 = Leaf then Leaf else
if t2 = Leaf then Leaf else
- case t1 of Node _ l1 k r1 \<Rightarrow>
+ case t1 of Node l1 k _ r1 \<Rightarrow>
let (l2,kin,r2) = split t2 k;
l' = inter l1 l2; r' = inter r1 r2
in if kin then join l' k r' else join2 l' r')"
@@ -196,7 +196,7 @@
proof (cases t1)
case Leaf thus ?thesis by (simp add: inter.simps)
next
- case [simp]: (Node _ l1 k r1)
+ case [simp]: (Node l1 k _ r1)
show ?thesis
proof (cases "t2 = Leaf")
case True thus ?thesis by (simp add: inter.simps)
@@ -246,7 +246,7 @@
"diff t1 t2 =
(if t1 = Leaf then Leaf else
if t2 = Leaf then t1 else
- case t2 of Node _ l2 k r2 \<Rightarrow>
+ case t2 of Node l2 k _ r2 \<Rightarrow>
let (l1,_,r1) = split t1 k;
l' = diff l1 l2; r' = diff r1 r2
in join2 l' r')"
@@ -261,7 +261,7 @@
proof (cases t2)
case Leaf thus ?thesis by (simp add: diff.simps)
next
- case [simp]: (Node _ l2 k r2)
+ case [simp]: (Node l2 k _ r2)
show ?thesis
proof (cases "t1 = Leaf")
case True thus ?thesis by (simp add: diff.simps)
@@ -341,7 +341,7 @@
end
interpretation unbal: Set2_Join
-where join = "\<lambda>l x r. Node () l x r" and inv = "\<lambda>t. True"
+where join = "\<lambda>l x r. Node l x () r" and inv = "\<lambda>t. True"
proof (standard, goal_cases)
case 1 show ?case by simp
next
--- a/src/HOL/Data_Structures/Tree2.thy Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/Tree2.thy Mon Jun 11 16:29:27 2018 +0200
@@ -4,30 +4,30 @@
datatype ('a,'b) tree =
Leaf ("\<langle>\<rangle>") |
- Node 'b "('a,'b)tree" 'a "('a,'b) tree" ("(1\<langle>_,/ _,/ _,/ _\<rangle>)")
+ Node "('a,'b)tree" 'a 'b "('a,'b) tree" ("(1\<langle>_,/ _,/ _,/ _\<rangle>)")
fun inorder :: "('a,'b)tree \<Rightarrow> 'a list" where
"inorder Leaf = []" |
-"inorder (Node _ l a r) = inorder l @ a # inorder r"
+"inorder (Node l a _ r) = inorder l @ a # inorder r"
fun height :: "('a,'b) tree \<Rightarrow> nat" where
"height Leaf = 0" |
-"height (Node _ l a r) = max (height l) (height r) + 1"
+"height (Node l a _ r) = max (height l) (height r) + 1"
fun set_tree :: "('a,'b) tree \<Rightarrow> 'a set" where
"set_tree Leaf = {}" |
-"set_tree (Node _ l x r) = Set.insert x (set_tree l \<union> set_tree r)"
+"set_tree (Node l a _ r) = Set.insert a (set_tree l \<union> set_tree r)"
fun bst :: "('a::linorder,'b) tree \<Rightarrow> bool" where
"bst Leaf = True" |
-"bst (Node _ l a r) = (bst l \<and> bst r \<and> (\<forall>x \<in> set_tree l. x < a) \<and> (\<forall>x \<in> set_tree r. a < x))"
+"bst (Node l a _ r) = (bst l \<and> bst r \<and> (\<forall>x \<in> set_tree l. x < a) \<and> (\<forall>x \<in> set_tree r. a < x))"
definition size1 :: "('a,'b) tree \<Rightarrow> nat" where
"size1 t = size t + 1"
lemma size1_simps[simp]:
"size1 \<langle>\<rangle> = 1"
- "size1 \<langle>u, l, x, r\<rangle> = size1 l + size1 r"
+ "size1 \<langle>l, x, u, r\<rangle> = size1 l + size1 r"
by (simp_all add: size1_def)
lemma size1_ge0[simp]: "0 < size1 t"