src/HOL/Data_Structures/Tree_Set.thy
changeset 71829 6f2663df8374
parent 71463 a31a9da43694
equal deleted inserted replaced
71828:415c38ef38c0 71829:6f2663df8374
     8   Cmp
     8   Cmp
     9   Set_Specs
     9   Set_Specs
    10 begin
    10 begin
    11 
    11 
    12 definition empty :: "'a tree" where
    12 definition empty :: "'a tree" where
    13 "empty == Leaf"
    13 "empty = Leaf"
    14 
    14 
    15 fun isin :: "'a::linorder tree \<Rightarrow> 'a \<Rightarrow> bool" where
    15 fun isin :: "'a::linorder tree \<Rightarrow> 'a \<Rightarrow> bool" where
    16 "isin Leaf x = False" |
    16 "isin Leaf x = False" |
    17 "isin (Node l a r) x =
    17 "isin (Node l a r) x =
    18   (case cmp x a of
    18   (case cmp x a of
    42   (case cmp x a of
    42   (case cmp x a of
    43      LT \<Rightarrow>  Node (delete x l) a r |
    43      LT \<Rightarrow>  Node (delete x l) a r |
    44      GT \<Rightarrow>  Node l a (delete x r) |
    44      GT \<Rightarrow>  Node l a (delete x r) |
    45      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 
    46 
    47 text \<open>Deletion by appending:\<close>
    47 text \<open>Deletion by joining:\<close>
    48 
    48 
    49 fun app :: "('a::linorder)tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
    49 fun join :: "('a::linorder)tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
    50 "app t Leaf = t" |
    50 "join t Leaf = t" |
    51 "app Leaf t = t" |
    51 "join Leaf t = t" |
    52 "app (Node t1 a t2) (Node t3 b t4) =
    52 "join (Node t1 a t2) (Node t3 b t4) =
    53   (case app t2 t3 of
    53   (case join t2 t3 of
    54      Leaf \<Rightarrow> Node t1 a (Node Leaf b t4) |
    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))"
    55      Node u2 x u3 \<Rightarrow> Node (Node t1 a u2) x (Node u3 b t4))"
    56 
    56 
    57 fun delete2 :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
    57 fun delete2 :: "'a::linorder \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where
    58 "delete2 x Leaf = Leaf" |
    58 "delete2 x Leaf = Leaf" |
    59 "delete2 x (Node l a r) =
    59 "delete2 x (Node l a r) =
    60   (case cmp x a of
    60   (case cmp x a of
    61      LT \<Rightarrow>  Node (delete2 x l) a r |
    61      LT \<Rightarrow>  Node (delete2 x l) a r |
    62      GT \<Rightarrow>  Node l a (delete2 x r) |
    62      GT \<Rightarrow>  Node l a (delete2 x r) |
    63      EQ \<Rightarrow> app l r)"
    63      EQ \<Rightarrow> join l r)"
    64 
    64 
    65 
    65 
    66 subsection "Functional Correctness Proofs"
    66 subsection "Functional Correctness Proofs"
    67 
    67 
    68 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))"
    93   case 3 thus ?case by(simp add: inorder_insert)
    93   case 3 thus ?case by(simp add: inorder_insert)
    94 next
    94 next
    95   case 4 thus ?case by(simp add: inorder_delete)
    95   case 4 thus ?case by(simp add: inorder_delete)
    96 qed (rule TrueI)+
    96 qed (rule TrueI)+
    97 
    97 
    98 lemma inorder_app:
    98 lemma inorder_join:
    99   "inorder(app l r) = inorder l @ inorder r"
    99   "inorder(join l r) = inorder l @ inorder r"
   100 by(induction l r rule: app.induct) (auto split: tree.split)
   100 by(induction l r rule: join.induct) (auto split: tree.split)
   101 
   101 
   102 lemma inorder_delete2:
   102 lemma inorder_delete2:
   103   "sorted(inorder t) \<Longrightarrow> inorder(delete2 x t) = del_list x (inorder t)"
   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)
   104 by(induction t) (auto simp: inorder_join del_list_simps)
   105 
   105 
   106 interpretation S2: Set_by_Ordered
   106 interpretation S2: Set_by_Ordered
   107 where empty = empty and isin = isin and insert = insert and delete = delete2
   107 where empty = empty and isin = isin and insert = insert and delete = delete2
   108 and inorder = inorder and inv = "\<lambda>_. True"
   108 and inorder = inorder and inv = "\<lambda>_. True"
   109 proof (standard, goal_cases)
   109 proof (standard, goal_cases)