| 
61640
 | 
     1  | 
(* Author: Tobias Nipkow *)
  | 
| 
 | 
     2  | 
  | 
| 
 | 
     3  | 
section {* Implementing Ordered Maps *}
 | 
| 
 | 
     4  | 
  | 
| 
 | 
     5  | 
theory Map_by_Ordered
  | 
| 
 | 
     6  | 
imports AList_Upd_Del
  | 
| 
 | 
     7  | 
begin
  | 
| 
 | 
     8  | 
  | 
| 
 | 
     9  | 
locale Map =
  | 
| 
 | 
    10  | 
fixes empty :: "'m"
  | 
| 
 | 
    11  | 
fixes update :: "'a \<Rightarrow> 'b \<Rightarrow> 'm \<Rightarrow> 'm"
  | 
| 
 | 
    12  | 
fixes delete :: "'a \<Rightarrow> 'm \<Rightarrow> 'm"
  | 
| 
62196
 | 
    13  | 
fixes lookup :: "'m \<Rightarrow> 'a \<Rightarrow> 'b option"
  | 
| 
61640
 | 
    14  | 
fixes invar :: "'m \<Rightarrow> bool"
  | 
| 
62196
 | 
    15  | 
assumes map_empty: "lookup empty = (\<lambda>_. None)"
  | 
| 
 | 
    16  | 
and map_update: "invar m \<Longrightarrow> lookup(update a b m) = (lookup m)(a := Some b)"
  | 
| 
 | 
    17  | 
and map_delete: "invar m \<Longrightarrow> lookup(delete a m) = (lookup m)(a := None)"
  | 
| 
61686
 | 
    18  | 
and invar_empty: "invar empty"
  | 
| 
 | 
    19  | 
and invar_update: "invar m \<Longrightarrow> invar(update a b m)"
  | 
| 
 | 
    20  | 
and invar_delete: "invar m \<Longrightarrow> invar(delete a m)"
  | 
| 
61640
 | 
    21  | 
  | 
| 
 | 
    22  | 
locale Map_by_Ordered =
  | 
| 
 | 
    23  | 
fixes empty :: "'t"
  | 
| 
 | 
    24  | 
fixes update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> 't \<Rightarrow> 't"
  | 
| 
 | 
    25  | 
fixes delete :: "'a \<Rightarrow> 't \<Rightarrow> 't"
  | 
| 
 | 
    26  | 
fixes lookup :: "'t \<Rightarrow> 'a \<Rightarrow> 'b option"
  | 
| 
 | 
    27  | 
fixes inorder :: "'t \<Rightarrow> ('a * 'b) list"
 | 
| 
61686
 | 
    28  | 
fixes inv :: "'t \<Rightarrow> bool"
  | 
| 
61640
 | 
    29  | 
assumes empty: "inorder empty = []"
  | 
| 
61686
 | 
    30  | 
and lookup: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
  | 
| 
61640
 | 
    31  | 
  lookup t a = map_of (inorder t) a"
  | 
| 
61686
 | 
    32  | 
and update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
  | 
| 
61640
 | 
    33  | 
  inorder(update a b t) = upd_list a b (inorder t)"
  | 
| 
61686
 | 
    34  | 
and delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
  | 
| 
61640
 | 
    35  | 
  inorder(delete a t) = del_list a (inorder t)"
  | 
| 
61686
 | 
    36  | 
and inv_empty:  "inv empty"
  | 
| 
61688
 | 
    37  | 
and inv_update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(update a b t)"
  | 
| 
61686
 | 
    38  | 
and inv_delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(delete a t)"
  | 
| 
61640
 | 
    39  | 
begin
  | 
| 
 | 
    40  | 
  | 
| 
 | 
    41  | 
sublocale Map
  | 
| 
61688
 | 
    42  | 
  empty update delete lookup "\<lambda>t. inv t \<and> sorted1 (inorder t)"
  | 
| 
61640
 | 
    43  | 
proof(standard, goal_cases)
  | 
| 
61688
 | 
    44  | 
  case 1 show ?case by (auto simp: lookup empty inv_empty)
  | 
| 
61640
 | 
    45  | 
next
  | 
| 
61688
 | 
    46  | 
  case 2 thus ?case
  | 
| 
 | 
    47  | 
    by(simp add: fun_eq_iff update inv_update map_of_ins_list lookup sorted_upd_list)
  | 
| 
61640
 | 
    48  | 
next
  | 
| 
61688
 | 
    49  | 
  case 3 thus ?case
  | 
| 
 | 
    50  | 
    by(simp add: fun_eq_iff delete inv_delete map_of_del_list lookup sorted_del_list)
  | 
| 
61640
 | 
    51  | 
next
  | 
| 
61686
 | 
    52  | 
  case 4 thus ?case by(simp add: empty inv_empty)
  | 
| 
61640
 | 
    53  | 
next
  | 
| 
61688
 | 
    54  | 
  case 5 thus ?case by(simp add: update inv_update sorted_upd_list)
  | 
| 
61640
 | 
    55  | 
next
  | 
| 
61686
 | 
    56  | 
  case 6 thus ?case by (auto simp: delete inv_delete sorted_del_list)
  | 
| 
61640
 | 
    57  | 
qed
  | 
| 
 | 
    58  | 
  | 
| 
 | 
    59  | 
end
  | 
| 
 | 
    60  | 
  | 
| 
 | 
    61  | 
end
  |