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