src/HOL/Data_Structures/Tree_Set.thy
changeset 71463 a31a9da43694
parent 68440 6826718f732d
child 71829 6f2663df8374
equal deleted inserted replaced
71462:ed8d50969995 71463:a31a9da43694
    28   (case cmp x a of
    28   (case cmp x a of
    29      LT \<Rightarrow> Node (insert x l) a r |
    29      LT \<Rightarrow> Node (insert x l) a r |
    30      EQ \<Rightarrow> Node l a r |
    30      EQ \<Rightarrow> Node l a r |
    31      GT \<Rightarrow> Node l a (insert x r))"
    31      GT \<Rightarrow> Node l a (insert x r))"
    32 
    32 
       
    33 text \<open>Deletion by replacing:\<close>
       
    34 
    33 fun split_min :: "'a tree \<Rightarrow> 'a * 'a tree" where
    35 fun split_min :: "'a tree \<Rightarrow> 'a * 'a tree" where
    34 "split_min (Node l a r) =
    36 "split_min (Node l a r) =
    35   (if l = Leaf then (a,r) else let (x,l') = split_min l in (x, Node l' a r))"
    37   (if l = Leaf then (a,r) else let (x,l') = split_min l in (x, Node l' a r))"
    36 
    38 
    37 fun delete :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
    39 fun delete :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
    39 "delete x (Node l a r) =
    41 "delete x (Node l a r) =
    40   (case cmp x a of
    42   (case cmp x a of
    41      LT \<Rightarrow>  Node (delete x l) a r |
    43      LT \<Rightarrow>  Node (delete x l) a r |
    42      GT \<Rightarrow>  Node l a (delete x r) |
    44      GT \<Rightarrow>  Node l a (delete x r) |
    43      EQ \<Rightarrow> if r = Leaf then l else let (a',r') = split_min r in Node l a' r')"
    45      EQ \<Rightarrow> if r = Leaf then l else let (a',r') = split_min r in Node l a' r')"
       
    46 
       
    47 text \<open>Deletion by appending:\<close>
       
    48 
       
    49 fun app :: "('a::linorder)tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
       
    50 "app t Leaf = t" |
       
    51 "app Leaf t = t" |
       
    52 "app (Node t1 a t2) (Node t3 b t4) =
       
    53   (case app t2 t3 of
       
    54      Leaf \<Rightarrow> Node t1 a (Node Leaf b t4) |
       
    55      Node u2 x u3 \<Rightarrow> Node (Node t1 a u2) x (Node u3 b t4))"
       
    56 
       
    57 fun delete2 :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
       
    58 "delete2 x Leaf = Leaf" |
       
    59 "delete2 x (Node l a r) =
       
    60   (case cmp x a of
       
    61      LT \<Rightarrow>  Node (delete2 x l) a r |
       
    62      GT \<Rightarrow>  Node l a (delete2 x r) |
       
    63      EQ \<Rightarrow> app l r)"
    44 
    64 
    45 
    65 
    46 subsection "Functional Correctness Proofs"
    66 subsection "Functional Correctness Proofs"
    47 
    67 
    48 lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> set (inorder t))"
    68 lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> set (inorder t))"
    73   case 3 thus ?case by(simp add: inorder_insert)
    93   case 3 thus ?case by(simp add: inorder_insert)
    74 next
    94 next
    75   case 4 thus ?case by(simp add: inorder_delete)
    95   case 4 thus ?case by(simp add: inorder_delete)
    76 qed (rule TrueI)+
    96 qed (rule TrueI)+
    77 
    97 
       
    98 lemma inorder_app:
       
    99   "inorder(app l r) = inorder l @ inorder r"
       
   100 by(induction l r rule: app.induct) (auto split: tree.split)
       
   101 
       
   102 lemma inorder_delete2:
       
   103   "sorted(inorder t) \<Longrightarrow> inorder(delete2 x t) = del_list x (inorder t)"
       
   104 by(induction t) (auto simp: inorder_app del_list_simps)
       
   105 
       
   106 interpretation S2: Set_by_Ordered
       
   107 where empty = empty and isin = isin and insert = insert and delete = delete2
       
   108 and inorder = inorder and inv = "\<lambda>_. True"
       
   109 proof (standard, goal_cases)
       
   110   case 1 show ?case by (simp add: empty_def)
       
   111 next
       
   112   case 2 thus ?case by(simp add: isin_set)
       
   113 next
       
   114   case 3 thus ?case by(simp add: inorder_insert)
       
   115 next
       
   116   case 4 thus ?case by(simp add: inorder_delete2)
       
   117 qed (rule TrueI)+
       
   118 
    78 end
   119 end