| 61203 |      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"
 | 
|  |     13 | fixes map_of :: "'m \<Rightarrow> 'a \<Rightarrow> 'b option"
 | 
|  |     14 | fixes invar :: "'m \<Rightarrow> bool"
 | 
|  |     15 | assumes "map_of empty = (\<lambda>_. None)"
 | 
|  |     16 | assumes "invar m \<Longrightarrow> map_of(update a b m) = (map_of m)(a := Some b)"
 | 
|  |     17 | assumes "invar m \<Longrightarrow> map_of(delete a m) = (map_of m)(a := None)"
 | 
| 61428 |     18 | assumes "invar empty"
 | 
| 61203 |     19 | assumes "invar m \<Longrightarrow> invar(update a b m)"
 | 
|  |     20 | assumes "invar m \<Longrightarrow> invar(delete a m)"
 | 
|  |     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"
 | 
|  |     28 | fixes wf :: "'t \<Rightarrow> bool"
 | 
|  |     29 | assumes empty: "inorder empty = []"
 | 
|  |     30 | assumes lookup: "wf t \<and> sorted1 (inorder t) \<Longrightarrow>
 | 
|  |     31 |   lookup t a = map_of (inorder t) a"
 | 
|  |     32 | assumes update: "wf t \<and> sorted1 (inorder t) \<Longrightarrow>
 | 
|  |     33 |   inorder(update a b t) = upd_list a b (inorder t)"
 | 
|  |     34 | assumes delete: "wf t \<and> sorted1 (inorder t) \<Longrightarrow>
 | 
|  |     35 |   inorder(delete a t) = del_list a (inorder t)"
 | 
| 61428 |     36 | assumes wf_empty:  "wf empty"
 | 
| 61203 |     37 | assumes wf_insert: "wf t \<and> sorted1 (inorder t) \<Longrightarrow> wf(update a b t)"
 | 
|  |     38 | assumes wf_delete: "wf t \<and> sorted1 (inorder t) \<Longrightarrow> wf(delete a t)"
 | 
|  |     39 | begin
 | 
|  |     40 | 
 | 
|  |     41 | sublocale Map
 | 
|  |     42 |   empty update delete "map_of o inorder" "\<lambda>t. wf t \<and> sorted1 (inorder t)"
 | 
|  |     43 | proof(standard, goal_cases)
 | 
|  |     44 |   case 1 show ?case by (auto simp: empty)
 | 
|  |     45 | next
 | 
|  |     46 |   case 2 thus ?case by(simp add: update map_of_ins_list)
 | 
|  |     47 | next
 | 
|  |     48 |   case 3 thus ?case by(simp add: delete map_of_del_list)
 | 
|  |     49 | next
 | 
| 61428 |     50 |   case 4 thus ?case by(simp add: empty wf_empty)
 | 
| 61203 |     51 | next
 | 
| 61428 |     52 |   case 5 thus ?case by(simp add: update wf_insert sorted_upd_list)
 | 
|  |     53 | next
 | 
|  |     54 |   case 6 thus ?case by (auto simp: delete wf_delete sorted_del_list)
 | 
| 61203 |     55 | qed
 | 
|  |     56 | 
 | 
|  |     57 | end
 | 
|  |     58 | 
 | 
|  |     59 | end
 |