src/HOL/Data_Structures/AA_Map.thy
changeset 62496 f187aaf602c4
parent 62130 90a3016a6c12
child 63411 e051eea34990
equal deleted inserted replaced
62495:83db706d7771 62496:f187aaf602c4
     1 (* Author: Tobias Nipkow *)
     1 (* Author: Tobias Nipkow *)
     2 
     2 
     3 section "AA Implementation of Maps"
     3 section "AA Tree Implementation of Maps"
     4 
     4 
     5 theory AA_Map
     5 theory AA_Map
     6 imports
     6 imports
     7   AA_Set
     7   AA_Set
     8   Lookup2
     8   Lookup2
    24      GT \<Rightarrow> adjust (Node lv l (a,b) (delete x r)) |
    24      GT \<Rightarrow> adjust (Node lv l (a,b) (delete x r)) |
    25      EQ \<Rightarrow> (if l = Leaf then r
    25      EQ \<Rightarrow> (if l = Leaf then r
    26             else let (l',ab') = del_max l in adjust (Node lv l' ab' r)))"
    26             else let (l',ab') = del_max l in adjust (Node lv l' ab' r)))"
    27 
    27 
    28 
    28 
       
    29 subsection "Invariance"
       
    30 
       
    31 subsubsection "Proofs for insert"
       
    32 
       
    33 lemma lvl_update_aux:
       
    34   "lvl (update x y t) = lvl t \<or> lvl (update x y t) = lvl t + 1 \<and> sngl (update x y t)"
       
    35 apply(induction t)
       
    36 apply (auto simp: lvl_skew)
       
    37 apply (metis Suc_eq_plus1 lvl.simps(2) lvl_split lvl_skew)+
       
    38 done
       
    39 
       
    40 lemma lvl_update: obtains
       
    41   (Same) "lvl (update x y t) = lvl t" |
       
    42   (Incr) "lvl (update x y t) = lvl t + 1" "sngl (update x y t)"
       
    43 using lvl_update_aux by fastforce
       
    44 
       
    45 declare invar.simps(2)[simp]
       
    46 
       
    47 lemma lvl_update_sngl: "invar t \<Longrightarrow> sngl t \<Longrightarrow> lvl(update x y t) = lvl t"
       
    48 proof (induction t rule: update.induct)
       
    49   case (2 x y lv t1 a b t2)
       
    50   consider (LT) "x < a" | (GT) "x > a" | (EQ) "x = a" 
       
    51     using less_linear by blast 
       
    52   thus ?case proof cases
       
    53     case LT
       
    54     thus ?thesis using 2 by (auto simp add: skew_case split_case split: tree.splits)
       
    55   next
       
    56     case GT
       
    57     thus ?thesis using 2 proof (cases t1)
       
    58       case Node
       
    59       thus ?thesis using 2 GT  
       
    60         apply (auto simp add: skew_case split_case split: tree.splits)
       
    61         by (metis less_not_refl2 lvl.simps(2) lvl_update_aux n_not_Suc_n sngl.simps(3))+ 
       
    62     qed (auto simp add: lvl_0_iff)
       
    63   qed simp
       
    64 qed simp
       
    65 
       
    66 lemma lvl_update_incr_iff: "(lvl(update a b t) = lvl t + 1) \<longleftrightarrow>
       
    67   (EX l x r. update a b t = Node (lvl t + 1) l x r \<and> lvl l = lvl r)"
       
    68 apply(cases t)
       
    69 apply(auto simp add: skew_case split_case split: if_splits)
       
    70 apply(auto split: tree.splits if_splits)
       
    71 done
       
    72 
       
    73 lemma invar_update: "invar t \<Longrightarrow> invar(update a b t)"
       
    74 proof(induction t)
       
    75   case (Node n l xy r)
       
    76   hence il: "invar l" and ir: "invar r" by auto
       
    77   obtain x y where [simp]: "xy = (x,y)" by fastforce
       
    78   note N = Node
       
    79   let ?t = "Node n l xy r"
       
    80   have "a < x \<or> a = x \<or> x < a" by auto
       
    81   moreover
       
    82   { assume "a < x"
       
    83     note iil = Node.IH(1)[OF il]
       
    84     have ?case
       
    85     proof (cases rule: lvl_update[of a b l])
       
    86       case (Same) thus ?thesis
       
    87         using \<open>a<x\<close> invar_NodeL[OF Node.prems iil Same]
       
    88         by (simp add: skew_invar split_invar del: invar.simps)
       
    89     next
       
    90       case (Incr)
       
    91       then obtain t1 w t2 where ial[simp]: "update a b l = Node n t1 w t2"
       
    92         using Node.prems by (auto simp: lvl_Suc_iff)
       
    93       have l12: "lvl t1 = lvl t2"
       
    94         by (metis Incr(1) ial lvl_update_incr_iff tree.inject)
       
    95       have "update a b ?t = split(skew(Node n (update a b l) xy r))"
       
    96         by(simp add: \<open>a<x\<close>)
       
    97       also have "skew(Node n (update a b l) xy r) = Node n t1 w (Node n t2 xy r)"
       
    98         by(simp)
       
    99       also have "invar(split \<dots>)"
       
   100       proof (cases r)
       
   101         case Leaf
       
   102         hence "l = Leaf" using Node.prems by(auto simp: lvl_0_iff)
       
   103         thus ?thesis using Leaf ial by simp
       
   104       next
       
   105         case [simp]: (Node m t3 y t4)
       
   106         show ?thesis (*using N(3) iil l12 by(auto)*)
       
   107         proof cases
       
   108           assume "m = n" thus ?thesis using N(3) iil by(auto)
       
   109         next
       
   110           assume "m \<noteq> n" thus ?thesis using N(3) iil l12 by(auto)
       
   111         qed
       
   112       qed
       
   113       finally show ?thesis .
       
   114     qed
       
   115   }
       
   116   moreover
       
   117   { assume "x < a"
       
   118     note iir = Node.IH(2)[OF ir]
       
   119     from \<open>invar ?t\<close> have "n = lvl r \<or> n = lvl r + 1" by auto
       
   120     hence ?case
       
   121     proof
       
   122       assume 0: "n = lvl r"
       
   123       have "update a b ?t = split(skew(Node n l xy (update a b r)))"
       
   124         using \<open>a>x\<close> by(auto)
       
   125       also have "skew(Node n l xy (update a b r)) = Node n l xy (update a b r)"
       
   126         using Node.prems by(simp add: skew_case split: tree.split)
       
   127       also have "invar(split \<dots>)"
       
   128       proof -
       
   129         from lvl_update_sngl[OF ir sngl_if_invar[OF \<open>invar ?t\<close> 0], of a b]
       
   130         obtain t1 p t2 where iar: "update a b r = Node n t1 p t2"
       
   131           using Node.prems 0 by (auto simp: lvl_Suc_iff)
       
   132         from Node.prems iar 0 iir
       
   133         show ?thesis by (auto simp: split_case split: tree.splits)
       
   134       qed
       
   135       finally show ?thesis .
       
   136     next
       
   137       assume 1: "n = lvl r + 1"
       
   138       hence "sngl ?t" by(cases r) auto
       
   139       show ?thesis
       
   140       proof (cases rule: lvl_update[of a b r])
       
   141         case (Same)
       
   142         show ?thesis using \<open>x<a\<close> il ir invar_NodeR[OF Node.prems 1 iir Same]
       
   143           by (auto simp add: skew_invar split_invar)
       
   144       next
       
   145         case (Incr)
       
   146         thus ?thesis using invar_NodeR2[OF `invar ?t` Incr(2) 1 iir] 1 \<open>x < a\<close>
       
   147           by (auto simp add: skew_invar split_invar split: if_splits)
       
   148       qed
       
   149     qed
       
   150   }
       
   151   moreover { assume "a = x" hence ?case using Node.prems by auto }
       
   152   ultimately show ?case by blast
       
   153 qed simp
       
   154 
       
   155 subsubsection "Proofs for delete"
       
   156 
       
   157 declare invar.simps(2)[simp del]
       
   158 
       
   159 theorem post_delete: "invar t \<Longrightarrow> post_del t (delete x t)"
       
   160 proof (induction t)
       
   161   case (Node lv l ab r)
       
   162 
       
   163   obtain a b where [simp]: "ab = (a,b)" by fastforce
       
   164 
       
   165   let ?l' = "delete x l" and ?r' = "delete x r"
       
   166   let ?t = "Node lv l ab r" let ?t' = "delete x ?t"
       
   167 
       
   168   from Node.prems have inv_l: "invar l" and inv_r: "invar r" by (auto)
       
   169 
       
   170   note post_l' = Node.IH(1)[OF inv_l]
       
   171   note preL = pre_adj_if_postL[OF Node.prems post_l']
       
   172 
       
   173   note post_r' = Node.IH(2)[OF inv_r]
       
   174   note preR = pre_adj_if_postR[OF Node.prems post_r']
       
   175 
       
   176   show ?case
       
   177   proof (cases rule: linorder_cases[of x a])
       
   178     case less
       
   179     thus ?thesis using Node.prems by (simp add: post_del_adjL preL)
       
   180   next
       
   181     case greater
       
   182     thus ?thesis using Node.prems preR by (simp add: post_del_adjR post_r')
       
   183   next
       
   184     case equal
       
   185     show ?thesis
       
   186     proof cases
       
   187       assume "l = Leaf" thus ?thesis using equal Node.prems
       
   188         by(auto simp: post_del_def invar.simps(2))
       
   189     next
       
   190       assume "l \<noteq> Leaf" thus ?thesis using equal Node.prems
       
   191         by simp (metis inv_l post_del_adjL post_del_max pre_adj_if_postL)
       
   192     qed
       
   193   qed
       
   194 qed (simp add: post_del_def)
       
   195 
       
   196 
    29 subsection {* Functional Correctness Proofs *}
   197 subsection {* Functional Correctness Proofs *}
    30 
   198 
    31 theorem inorder_update:
   199 theorem inorder_update:
    32   "sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
   200   "sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
    33 by (induct t) (auto simp: upd_list_simps inorder_split inorder_skew)
   201 by (induct t) (auto simp: upd_list_simps inorder_split inorder_skew)
    34 
   202 
    35 
       
    36 theorem inorder_delete:
   203 theorem inorder_delete:
    37   "sorted1(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
   204   "\<lbrakk>invar t; sorted1(inorder t)\<rbrakk> \<Longrightarrow>
       
   205   inorder (delete x t) = del_list x (inorder t)"
    38 by(induction t)
   206 by(induction t)
    39   (auto simp: del_list_simps inorder_adjust del_maxD split: prod.splits)
   207   (auto simp: del_list_simps inorder_adjust pre_adj_if_postL pre_adj_if_postR 
    40 
   208               post_del_max post_delete del_maxD split: prod.splits)
    41 interpretation Map_by_Ordered
   209 
       
   210 interpretation I: Map_by_Ordered
    42 where empty = Leaf and lookup = lookup and update = update and delete = delete
   211 where empty = Leaf and lookup = lookup and update = update and delete = delete
    43 and inorder = inorder and inv = "\<lambda>_. True"
   212 and inorder = inorder and inv = invar
    44 proof (standard, goal_cases)
   213 proof (standard, goal_cases)
    45   case 1 show ?case by simp
   214   case 1 show ?case by simp
    46 next
   215 next
    47   case 2 thus ?case by(simp add: lookup_map_of)
   216   case 2 thus ?case by(simp add: lookup_map_of)
    48 next
   217 next
    49   case 3 thus ?case by(simp add: inorder_update)
   218   case 3 thus ?case by(simp add: inorder_update)
    50 next
   219 next
    51   case 4 thus ?case by(simp add: inorder_delete)
   220   case 4 thus ?case by(simp add: inorder_delete)
    52 qed auto
   221 next
       
   222   case 5 thus ?case by(simp)
       
   223 next
       
   224   case 6 thus ?case by(simp add: invar_update)
       
   225 next
       
   226   case 7 thus ?case using post_delete by(auto simp: post_del_def)
       
   227 qed
    53 
   228 
    54 end
   229 end