merged
authornipkow
Wed, 25 Sep 2019 18:39:47 +0200
changeset 70756 36283f21635c
parent 70754 59341cfb8fed (current diff)
parent 70755 3fb16bed5d6c (diff)
child 70757 6a835635fa93
merged
--- a/src/HOL/Data_Structures/AA_Map.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/AA_Map.thy	Wed Sep 25 18:39:47 2019 +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 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 \<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)"
+     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 l (a,b) lv r) =
+"delete x (Node l ((a,b), lv) r) =
   (case cmp x a of
-     LT \<Rightarrow> adjust (Node (delete x l) (a,b) lv r) |
-     GT \<Rightarrow> adjust (Node l (a,b) lv (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 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) \<longleftrightarrow>
-  (\<exists>l x r. update a b t = Node l x (lvl t + 1) 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)
 done
 
 lemma invar_update: "invar t \<Longrightarrow> 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 \<or> a = x \<or> 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: \<open>a<x\<close>)
-    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 \<dots>)"
-    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 \<open>a>x\<close> 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 \<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 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 \<Longrightarrow> 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)
 
--- a/src/HOL/Data_Structures/AA_Set.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/AA_Set.thy	Wed Sep 25 18:39:47 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 \<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 l a h 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 lr b h 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 (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 \<Rightarrow> '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 \<and> lvb = lvc \<comment> \<open>\<open>lva = lvc\<close> suffices\<close>
-    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 \<Rightarrow> 'a aa_tree \<Rightarrow> '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 \<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)"
+     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 _ _ Leaf) = True" |
+"sngl (Node _ (_, lva) (Node _ (_, lvb) _)) = (lva > lvb)"
 
 definition adjust :: "'a aa_tree \<Rightarrow> 'a aa_tree" where
 "adjust t =
  (case t of
-  Node l x lv 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 l x (lv-1) 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 t1 a lva (Node t2 b lvb t3)
-             \<Rightarrow> Node (Node t1 a lva t2) b (lvb+1) (Node t3 x (lv-1) 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 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 \<Rightarrow>
+        Node t1 (b,lvb) t4 \<Rightarrow>
           (case t1 of
-             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)))))"
+             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>\<open>adjust\<close> is expressed with the help of an
 incorrect auxiliary function \texttt{nlvl}.
@@ -81,20 +81,20 @@
 is not restored.\<close>
 
 fun split_max :: "'a aa_tree \<Rightarrow> '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 \<Rightarrow> 'a aa_tree \<Rightarrow> '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 \<Rightarrow> adjust (Node (delete x l) a lv r) |
-     GT \<Rightarrow> adjust (Node l a lv (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 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 \<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))))"
 
@@ -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)) \<Rightarrow>
+  Node t1 (x,lvx) (Node t2 (y,lvy) (Node t3 (z,lvz) t4)) \<Rightarrow>
    (if lvx = lvy \<and> 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 \<Rightarrow> t)"
 by(auto split: tree.split)
 
 lemma skew_case: "skew t = (case t of
-  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)
+  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 l a (Suc n) 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"
@@ -128,16 +128,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 l x lv (Node rl rx rlv rr)) =
-     (invar l \<and> invar \<langle>rl, rx, rlv, 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 l x lv 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 l a n 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+
 
 
@@ -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:
-  "\<lbrakk> invar(Node l x n r); invar l'; lvl l' = lvl l \<rbrakk> \<Longrightarrow> invar(Node l' x n 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 l x n r); n = lvl r + 1; invar r'; lvl r' = lvl r \<rbrakk> \<Longrightarrow> invar(Node l x n 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 l x n r); sngl r'; n = lvl r + 1; invar r'; lvl r' = n \<rbrakk> \<Longrightarrow> invar(Node l x n 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 l x (lvl t + 1) r \<and> lvl l = lvl r)"
-apply(cases t)
+  (\<exists>l x r. insert a t = Node l (x, lvl t + 1) r \<and> 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 \<Longrightarrow> 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 \<or> a = x \<or> 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: \<open>a<x\<close>)
-    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 \<dots>)"
-    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 \<open>a>x\<close> 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 \<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 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 \<langle>l, a, lv, 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"
+lemma invarR: "ASSUMPTION(invar \<langle>l, (a,lv), r\<rangle>) \<Longrightarrow> invar r"
 by(simp add: ASSUMPTION_def)
 
 lemma sngl_NodeI:
-  "sngl (Node l a lv r) \<Longrightarrow> sngl (Node l' a' lv r)"
-by(cases r) (simp_all)
+  "sngl (Node l (a,lv) r) \<Longrightarrow> 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 \<and> invar r \<and>
     lv = Suc (lvl r) \<and> 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)) \<or> 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)) \<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>l, a, lvl r, 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 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>)" 
+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
@@ -361,38 +363,38 @@
   (lvl t' = lvl t \<and> sngl t \<longrightarrow> sngl t')"
 
 lemma pre_adj_if_postR:
-  "invar\<langle>lv, l, a, r\<rangle> \<Longrightarrow> post_del r r' \<Longrightarrow> pre_adjust \<langle>lv, l, a, r'\<rangle>"
+  "invar\<langle>lv, (l, a), r\<rangle> \<Longrightarrow> post_del r r' \<Longrightarrow> pre_adjust \<langle>lv, (l, a), r'\<rangle>"
 by(cases "sngl r")
   (auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims)
 
 lemma pre_adj_if_postL:
-  "invar\<langle>l, a, lv, r\<rangle> \<Longrightarrow> post_del l l' \<Longrightarrow> pre_adjust \<langle>l', b, lv, 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>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>)"
+  "\<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))
 
 lemma post_del_adjR:
-assumes "invar\<langle>lv, l, a, r\<rangle>" "pre_adjust \<langle>lv, l, a, r'\<rangle>" "post_del r r'"
-shows "post_del \<langle>lv, l, a, r\<rangle> (adjust \<langle>lv, l, a, r'\<rangle>)"
+assumes "invar\<langle>l, (a,lv), r\<rangle>" "pre_adjust \<langle>l, (a,lv), r'\<rangle>" "post_del r r'"
+shows "post_del \<langle>l, (a,lv), r\<rangle> (adjust \<langle>l, (a,lv), r'\<rangle>)"
 proof(unfold post_del_def, safe del: disjCI)
-  let ?t = "\<langle>lv, l, a, r\<rangle>"
-  let ?t' = "adjust \<langle>lv, l, a, r'\<rangle>"
+  let ?t = "\<langle>l, (a,lv), r\<rangle>"
+  let ?t' = "adjust \<langle>l, (a,lv), r'\<rangle>"
   show "invar ?t'" by(rule invar_adjust[OF assms(2)])
   show "lvl ?t' = lvl ?t \<or> 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 \<langle>lv, l, a, r'\<rangle>"
-    proof(cases r')
+    have s: "sngl \<langle>l, (a,lv), r'\<rangle>"
+    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:
  "\<lbrakk> invar t; (t', x) = split_max t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow> post_del t t'"
 proof (induction t arbitrary: t' rule: split_max.induct)
-  case (2 lv l a lvr rl ra rr)
-  let ?r =  "\<langle>lvr, rl, ra, rr\<rangle>"
-  let ?t = "\<langle>lv, l, a, ?r\<rangle>"
+  case (2 l a lv rl bl rr)
+  let ?r =  "\<langle>rl, bl, rr\<rangle>"
+  let ?t = "\<langle>l, (a, lv), ?r\<rangle>"
   from "2.prems"(2) obtain r' where r': "(r', x) = split_max ?r"
-    and [simp]: "t' = adjust \<langle>lv, l, a, r'\<rangle>" by auto
+    and [simp]: "t' = adjust \<langle>l, (a, lv), r'\<rangle>" by auto
   from  "2.IH"[OF _ r'] \<open>invar ?t\<close> have post: "post_del ?r r'" by simp
   note preR = pre_adj_if_postR[OF \<open>invar ?t\<close> 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 \<Longrightarrow> 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)
 
--- a/src/HOL/Data_Structures/AVL_Map.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/AVL_Map.thy	Wed Sep 25 18:39:47 2019 +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 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 |
+"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 l (a,b) h r) = (case cmp x a of
-   EQ \<Rightarrow> del_root (Node l (a,b) h 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))"
 
@@ -114,7 +114,7 @@
   assumes "avl t" 
   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)
+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))
-      \<or> 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))
+      \<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
--- a/src/HOL/Data_Structures/AVL_Set.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/AVL_Set.thy	Wed Sep 25 18:39:47 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 \<Rightarrow> bool" where
 "avl Leaf = True" |
-"avl (Node l a h 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 l a h 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 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 \<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)"
 
@@ -47,38 +47,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 Leaf x 1 Leaf" |
-"insert x (Node l a h r) = (case cmp x a of
-   EQ \<Rightarrow> 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 \<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 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 \<Rightarrow> 'a avl_tree \<Rightarrow> '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 \<Rightarrow> del_root (Node l a h 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))"
 
@@ -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 \<Longrightarrow> ht t = height t"
-by (cases t) simp_all
+by (cases t rule: tree2_cases) simp_all
 
 lemma height_balL:
   "\<lbrakk> height l = height r + 2; avl l; avl r \<rbrakk> \<Longrightarrow>
@@ -171,7 +171,7 @@
   assumes "avl l" "avl r" and "height l = height r \<or> height l = height r + 1
     \<or> height r = height l + 1 \<or> 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 \<or> height l = height r + 1
     \<or> height r = height l + 1 \<or> 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 \<or> 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 \<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
@@ -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 \<open>avl t\<close> 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)) \<or> 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))
-      \<or> 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))
+      \<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
@@ -449,7 +449,7 @@
 
 lemma height_invers: 
   "(height t = 0) = (t = Leaf)"
-  "avl t \<Longrightarrow> (height t = Suc h) = (\<exists> l a r . t = Node l a (Suc h) 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>
@@ -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 \<and> height l = Suc h
     \<or> height r = Suc h \<and> height l = h
--- a/src/HOL/Data_Structures/Isin2.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/Isin2.thy	Wed Sep 25 18:39:47 2019 +0200
@@ -9,18 +9,18 @@
   Set_Specs
 begin
 
-fun isin :: "('a::linorder,'b) tree \<Rightarrow> 'a \<Rightarrow> bool" where
+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 |
      GT \<Rightarrow> isin r x)"
 
 lemma isin_set_inorder: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> 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 \<Longrightarrow> isin t x \<longleftrightarrow> x \<in> set_tree t"
-by(induction t) auto
+by(induction t rule: tree2_induct) auto
 
 end
--- a/src/HOL/Data_Structures/Leftist_Heap.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/Leftist_Heap.thy	Wed Sep 25 18:39:47 2019 +0200
@@ -10,30 +10,30 @@
   Complex_Main
 begin
 
-fun mset_tree :: "('a,'b) tree \<Rightarrow> 'a multiset" where
+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"
+type_synonym 'a lheap = "('a*nat)tree"
 
 fun rank :: "'a lheap \<Rightarrow> nat" where
 "rank Leaf = 0" |
-"rank (Node _ _ _ r) = rank r + 1"
+"rank (Node _ _ r) = rank r + 1"
 
 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
+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_tree l \<union> set_tree r. m \<le> x))"
 
 fun ltree :: "'a lheap \<Rightarrow> bool" where
 "ltree Leaf = True" |
-"ltree (Node l a n r) =
+"ltree (Node l (a, n) r) =
  (n = rank r + 1 \<and> rank l \<ge> rank r \<and> ltree l & ltree r)"
 
 definition empty :: "'a lheap" where
@@ -42,10 +42,10 @@
 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 l a (rr+1) r else Node r a (rl+1) 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 l a n r) = a"
+"get_min(Node l (a, n) r) = a"
 
 text \<open>For function \<open>merge\<close>:\<close>
 unbundle pattern_aliases
@@ -53,7 +53,7 @@
 fun merge :: "'a::ord lheap \<Rightarrow> 'a lheap \<Rightarrow> '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 \<le> 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, _) \<Rightarrow> t2 |
   (_, Leaf) \<Rightarrow> t1 |
-  (Node l1 a1 n1 r1, Node l2 a2 n2 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 t1 r2))"
 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 Leaf x 1 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 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 \<Longrightarrow> 2 ^ rank t \<le> 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 \<le> rank l" by simp
   hence *: "(2::nat) ^ rank r \<le> 2 ^ rank l" by simp
-  have "(2::nat) ^ rank \<langle>l, a, n, 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>l, a, n, r\<rangle>" by simp
+  also have "\<dots> = size1 \<langle>l, (a, n), r\<rangle>" by simp
   finally show ?case .
 qed
 
@@ -196,16 +196,16 @@
 fun t_merge :: "'a::ord lheap \<Rightarrow> 'a lheap \<Rightarrow> 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 \<le> a2 then 1 + t_merge r1 t2
    else 1 + t_merge t1 r2)"
 
 definition t_insert :: "'a::ord \<Rightarrow> 'a lheap \<Rightarrow> 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 \<Rightarrow> 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 \<le> rank l + rank r + 1"
 proof(induction l r rule: merge.induct)
@@ -219,7 +219,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 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 \<le> 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)
--- a/src/HOL/Data_Structures/Lookup2.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/Lookup2.thy	Wed Sep 25 18:39:47 2019 +0200
@@ -9,13 +9,13 @@
   Map_Specs
 begin
 
-fun lookup :: "('a::linorder * 'b, 'c) tree \<Rightarrow> 'a \<Rightarrow> 'b option" where
+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:
   "sorted1(inorder t) \<Longrightarrow> 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
--- a/src/HOL/Data_Structures/RBT.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/RBT.thy	Wed Sep 25 18:39:47 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 \<equiv> Node l a Red r"
-abbreviation B where "B l a r \<equiv> Node l a Black 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 a t2) b t3) c t4 = R (B t1 a t2) b (B t3 c 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 l a c 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 a t2) b t3 = R (B t1 a t2) b t3" |
--- a/src/HOL/Data_Structures/RBT_Map.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/RBT_Map.thy	Wed Sep 25 18:39:47 2019 +0200
@@ -24,7 +24,7 @@
 
 fun del :: "'a::linorder \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('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 \<Rightarrow> if l \<noteq> Leaf \<and> color l = Black
            then baldL (del x l) (a,b) r else R (del x l) (a,b) r |
      GT \<Rightarrow> if r \<noteq> Leaf\<and> color r = Black
--- a/src/HOL/Data_Structures/RBT_Set.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/RBT_Set.thy	Wed Sep 25 18:39:47 2019 +0200
@@ -31,11 +31,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 |
@@ -100,11 +100,11 @@
 
 fun bheight :: "'a rbt \<Rightarrow> 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 \<Rightarrow> bool" where
 "invc Leaf = True" |
-"invc (Node l a c r) =
+"invc (Node l (a,c) r) =
   (invc l \<and> invc r \<and> (c = Red \<longrightarrow> color l = Black \<and> color r = Black))"
 
 text \<open>Weaker version:\<close>
@@ -113,10 +113,10 @@
 
 fun invh :: "'a rbt \<Rightarrow> bool" where
 "invh Leaf = True" |
-"invh (Node l x c 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+
+by (cases t rule: tree2_cases) simp+
 
 definition rbt :: "'a rbt \<Rightarrow> bool" where
 "rbt t = (invc t \<and> invh t \<and> 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 \<noteq> Leaf \<Longrightarrow> \<exists>c l x r. t = Node c l x r"
-by(cases t) auto
+lemma neq_LeafD: "t \<noteq> Leaf \<Longrightarrow> \<exists>l x c r. t = Node l (x,c) r"
+by(cases t rule: tree2_cases) auto
 
 lemma del_invc_invh: "invh t \<Longrightarrow> invc t \<Longrightarrow> invh (del x t) \<and>
    (color t = Red \<and> bheight (del x t) = bheight t \<and> invc (del x t) \<or>
--- a/src/HOL/Data_Structures/Set2_Join.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/Set2_Join.thy	Wed Sep 25 18:39:47 2019 +0200
@@ -20,26 +20,26 @@
 and recursion operators on it.\<close>
 
 locale Set2_Join =
-fixes join :: "('a::linorder,'b) tree \<Rightarrow> 'a \<Rightarrow> ('a,'b) tree \<Rightarrow> ('a,'b) tree"
-fixes inv :: "('a,'b) tree \<Rightarrow> bool"
+fixes join :: "('a::linorder*'b) tree \<Rightarrow> 'a \<Rightarrow> ('a*'b) tree \<Rightarrow> ('a*'b) tree"
+fixes inv :: "('a*'b) tree \<Rightarrow> bool"
 assumes set_join: "set_tree (join l a r) = set_tree l \<union> {a} \<union> set_tree r"
-assumes bst_join: "bst (Node l a b r) \<Longrightarrow> bst (join l a r)"
+assumes bst_join: "bst (Node l (a, b) r) \<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 a r)"
-assumes inv_Node: "\<lbrakk> inv (Node l a b r) \<rbrakk> \<Longrightarrow> inv l \<and> inv r"
+assumes inv_Node: "\<lbrakk> inv (Node l (a,b) r) \<rbrakk> \<Longrightarrow> inv l \<and> inv r"
 begin
 
 declare set_join [simp]
 
 subsection "\<open>split_min\<close>"
 
-fun split_min :: "('a,'b) tree \<Rightarrow> 'a \<times> ('a,'b) tree" where
-"split_min (Node l a _ r) =
+fun split_min :: "('a*'b) tree \<Rightarrow> 'a \<times> ('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:
   "\<lbrakk> split_min t = (m,t');  t \<noteq> Leaf \<rbrakk> \<Longrightarrow> m \<in> set_tree t \<and> set_tree t = {m} \<union> 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:
   "\<lbrakk> split_min t = (m,t');  bst t;  t \<noteq> Leaf \<rbrakk> \<Longrightarrow>  bst t' \<and> (\<forall>x \<in> 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:
   "\<lbrakk> split_min t = (m,t');  inv t;  t \<noteq> Leaf \<rbrakk> \<Longrightarrow>  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 "\<open>join2\<close>"
 
-definition join2 :: "('a,'b) tree \<Rightarrow> ('a,'b) tree \<Rightarrow> ('a,'b) tree" where
+definition join2 :: "('a*'b) tree \<Rightarrow> ('a*'b) tree \<Rightarrow> ('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 \<union> set_tree r"
@@ -80,9 +80,9 @@
 
 subsection "\<open>split\<close>"
 
-fun split :: "('a,'b)tree \<Rightarrow> 'a \<Rightarrow> ('a,'b)tree \<times> bool \<times> ('a,'b)tree" where
+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) x =
+"split (Node l (a, _) r) x =
   (case cmp x a of
      LT \<Rightarrow> let (l1,b,l2) = split l x in (l1, b, join l2 a r) |
      GT \<Rightarrow> 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) \<Longrightarrow> bst t \<Longrightarrow>
   set_tree l = {a \<in> set_tree t. a < x} \<and> set_tree r = {a \<in> set_tree t. x < a}
   \<and> (xin = (x \<in> set_tree t)) \<and> bst l \<and> 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) \<Longrightarrow> inv t \<Longrightarrow> inv l \<and> 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 "\<open>insert\<close>"
 
-definition insert :: "'a \<Rightarrow> ('a,'b) tree \<Rightarrow> ('a,'b) tree" where
+definition insert :: "'a \<Rightarrow> ('a*'b) tree \<Rightarrow> ('a*'b) tree" where
 "insert x t = (let (l,_,r) = split t x in join l x r)"
 
 lemma set_tree_insert: "bst t \<Longrightarrow> set_tree (insert x t) = {x} \<union> set_tree t"
@@ -125,7 +125,7 @@
 
 subsection "\<open>delete\<close>"
 
-definition delete :: "'a \<Rightarrow> ('a,'b) tree \<Rightarrow> ('a,'b) tree" where
+definition delete :: "'a \<Rightarrow> ('a*'b) tree \<Rightarrow> ('a*'b) tree" where
 "delete x t = (let (l,_,r) = split t x in join2 l r)"
 
 lemma set_tree_delete: "bst t \<Longrightarrow> set_tree (delete x t) = set_tree t - {x}"
@@ -140,11 +140,11 @@
 
 subsection "\<open>union\<close>"
 
-fun union :: "('a,'b)tree \<Rightarrow> ('a,'b)tree \<Rightarrow> ('a,'b)tree" where
+fun union :: "('a*'b)tree \<Rightarrow> ('a*'b)tree \<Rightarrow> ('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 \<Rightarrow>
+   case t1 of Node l1 (a, _) r1 \<Rightarrow>
    let (l2,_ ,r2) = split t2 a;
        l' = union l1 l2; r' = union r1 r2
    in join l' a r')"
@@ -176,11 +176,11 @@
 
 subsection "\<open>inter\<close>"
 
-fun inter :: "('a,'b)tree \<Rightarrow> ('a,'b)tree \<Rightarrow> ('a,'b)tree" where
+fun inter :: "('a*'b)tree \<Rightarrow> ('a*'b)tree \<Rightarrow> ('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 \<Rightarrow>
+   case t1 of Node l1 (a, _) r1 \<Rightarrow>
    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 \<inter> ?R2 = {}" "a \<notin> ?L2 \<union> ?R2" "?L1 \<inter> ?R2 = {}" "?L2 \<inter> ?R1 = {}"
         using split[OF sp] \<open>bst t1\<close> \<open>bst t2\<close> by (force, force, force, force, force)
       have IHl: "set_tree (inter l1 l2) = set_tree l1 \<inter> 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 \<inter> 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 \<inter> set_tree t2 = (?L1 \<union> ?R1 \<union> {a}) \<inter> (?L2 \<union> ?R2 \<union> ?K)"
         by(simp add: t2)
       also have "\<dots> = (?L1 \<inter> ?L2) \<union> (?R1 \<inter> ?R2) \<union> ?K"
@@ -241,11 +241,11 @@
 
 subsection "\<open>diff\<close>"
 
-fun diff :: "('a,'b)tree \<Rightarrow> ('a,'b)tree \<Rightarrow> ('a,'b)tree" where
+fun diff :: "('a*'b)tree \<Rightarrow> ('a*'b)tree \<Rightarrow> ('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 \<Rightarrow>
+   case t2 of Node l2 (a, _) r2 \<Rightarrow>
    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 \<notin> ?L1 \<union> ?R1" "?L1 \<inter> ?R2 = {}" "?L2 \<inter> ?R1 = {}"
         using split[OF sp] \<open>bst t1\<close> \<open>bst t2\<close> 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 \<union> ?R1) - (?L2 \<union> ?R2  \<union> {a})"
         by(simp add: t1)
       also have "\<dots> = (?L1 - ?L2) \<union> (?R1 - ?R2)"
@@ -340,7 +340,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/Set2_Join_RBT.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/Set2_Join_RBT.thy	Wed Sep 25 18:39:47 2019 +0200
@@ -179,7 +179,7 @@
 by(cases "(l,a,r)" rule: baliR.cases) (auto simp: ball_Un)
 
 lemma bst_joinL:
-  "\<lbrakk>bst (Node l a n r); bheight l \<le> bheight r\<rbrakk>
+  "\<lbrakk>bst (Node l (a, n) r); bheight l \<le> bheight r\<rbrakk>
   \<Longrightarrow> 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) \<Longrightarrow> bst (join l a r)"
+  "bst (Node l (a, n) r) \<Longrightarrow> bst (join l a r)"
 by(auto simp: bst_paint bst_joinL bst_joinR join_def)
 
 lemma inv_join: "\<lbrakk> invc l; invh l; invc r; invh r \<rbrakk> \<Longrightarrow> invc(join l x r) \<and> invh(join l x r)"
--- a/src/HOL/Data_Structures/Tree2.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/Tree2.thy	Wed Sep 25 18:39:47 2019 +0200
@@ -1,40 +1,29 @@
 theory Tree2
-imports Main
+imports "HOL-Library.Tree"
 begin
 
-datatype ('a,'b) tree =
-  Leaf ("\<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"
+text \<open>This theory provides the basic infrastructure for the type @{typ \<open>('a * 'b) tree\<close>}
+of augmented trees where @{typ 'a} is the key and @{typ 'b} some additional information.\<close>
 
-fun height :: "('a,'b) tree \<Rightarrow> nat" where
-"height Leaf = 0" |
-"height (Node l a _ r) = max (height l) (height r) + 1"
+text \<open>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"}.\<close>
 
-fun set_tree :: "('a,'b) tree \<Rightarrow> 'a set" where
-"set_tree Leaf = {}" |
-"set_tree (Node l a _ r) = Set.insert a (set_tree l \<union> 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 \<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))"
-
-fun size1 :: "('a,'b) tree \<Rightarrow> nat" where
-"size1 \<langle>\<rangle> = 1" |
-"size1 \<langle>l, _, _, r\<rangle> = size1 l + size1 r"
+fun inorder :: "('a*'b)tree \<Rightarrow> 'a list" where
+"inorder Leaf = []" |
+"inorder (Node l (a,_) r) = inorder l @ a # inorder r"
 
-fun complete :: "('a,'b) tree \<Rightarrow> bool" where
-"complete Leaf = True" |
-"complete (Node l _ _ r) = (complete l \<and> complete r \<and> height l = height r)"
+fun set_tree :: "('a*'b) tree \<Rightarrow> 'a set" where
+"set_tree Leaf = {}" |
+"set_tree (Node l (a,_) r) = Set.insert a (set_tree l \<union> 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 \<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))"
 
 lemma finite_set_tree[simp]: "finite(set_tree t)"
 by(induction t) auto
--- a/src/HOL/Data_Structures/Trie_Map.thy	Wed Sep 25 14:31:00 2019 +0200
+++ b/src/HOL/Data_Structures/Trie_Map.thy	Wed Sep 25 18:39:47 2019 +0200
@@ -21,10 +21,10 @@
 
 However, the development below works verbatim for any map implementation, eg \<open>Tree_Map\<close>,
 and not just \<open>RBT_Map\<close>, except for the termination lemma \<open>lookup_size\<close>.\<close>
-
+term size_tree
 lemma lookup_size[termination_simp]:
   fixes t :: "('a::linorder * 'a trie_map) rbt"
-  shows "lookup t a = Some b \<Longrightarrow> size b < Suc (size_tree (\<lambda>ab. Suc (size (snd ab))) (\<lambda>x. 0) t)"
+  shows "lookup t a = Some b \<Longrightarrow> size b < Suc (size_tree (\<lambda>ab. Suc(Suc (size (snd(fst ab))))) t)"
 apply(induction t a rule: lookup.induct)
 apply(auto split: if_splits)
 done