src/HOL/Data_Structures/Map_by_Ordered.thy
author nipkow
Mon, 11 Jan 2016 20:51:13 +0100
changeset 62130 90a3016a6c12
parent 61688 d04b1b4fb015
child 62196 66fb3d1767f2
permissions -rw-r--r--
added AA_Map; tuned titles
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
     1
(* Author: Tobias Nipkow *)
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
     2
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
     3
section {* Implementing Ordered Maps *}
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
     4
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
     5
theory Map_by_Ordered
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
     6
imports AList_Upd_Del
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
     7
begin
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
     8
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
     9
locale Map =
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    10
fixes empty :: "'m"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    11
fixes update :: "'a \<Rightarrow> 'b \<Rightarrow> 'm \<Rightarrow> 'm"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    12
fixes delete :: "'a \<Rightarrow> 'm \<Rightarrow> 'm"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    13
fixes map_of :: "'m \<Rightarrow> 'a \<Rightarrow> 'b option"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    14
fixes invar :: "'m \<Rightarrow> bool"
61686
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    15
assumes map_empty: "map_of empty = (\<lambda>_. None)"
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    16
and map_update: "invar m \<Longrightarrow> map_of(update a b m) = (map_of m)(a := Some b)"
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    17
and map_delete: "invar m \<Longrightarrow> map_of(delete a m) = (map_of m)(a := None)"
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    18
and invar_empty: "invar empty"
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    19
and invar_update: "invar m \<Longrightarrow> invar(update a b m)"
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    20
and invar_delete: "invar m \<Longrightarrow> invar(delete a m)"
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    21
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    22
locale Map_by_Ordered =
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    23
fixes empty :: "'t"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    24
fixes update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> 't \<Rightarrow> 't"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    25
fixes delete :: "'a \<Rightarrow> 't \<Rightarrow> 't"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    26
fixes lookup :: "'t \<Rightarrow> 'a \<Rightarrow> 'b option"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    27
fixes inorder :: "'t \<Rightarrow> ('a * 'b) list"
61686
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    28
fixes inv :: "'t \<Rightarrow> bool"
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    29
assumes empty: "inorder empty = []"
61686
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    30
and lookup: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    31
  lookup t a = map_of (inorder t) a"
61686
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    32
and update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    33
  inorder(update a b t) = upd_list a b (inorder t)"
61686
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    34
and delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    35
  inorder(delete a t) = del_list a (inorder t)"
61686
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    36
and inv_empty:  "inv empty"
61688
d04b1b4fb015 corrected inefficient implementation
nipkow
parents: 61686
diff changeset
    37
and inv_update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(update a b t)"
61686
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    38
and inv_delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(delete a t)"
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    39
begin
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    40
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    41
sublocale Map
61688
d04b1b4fb015 corrected inefficient implementation
nipkow
parents: 61686
diff changeset
    42
  empty update delete lookup "\<lambda>t. inv t \<and> sorted1 (inorder t)"
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    43
proof(standard, goal_cases)
61688
d04b1b4fb015 corrected inefficient implementation
nipkow
parents: 61686
diff changeset
    44
  case 1 show ?case by (auto simp: lookup empty inv_empty)
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    45
next
61688
d04b1b4fb015 corrected inefficient implementation
nipkow
parents: 61686
diff changeset
    46
  case 2 thus ?case
d04b1b4fb015 corrected inefficient implementation
nipkow
parents: 61686
diff changeset
    47
    by(simp add: fun_eq_iff update inv_update map_of_ins_list lookup sorted_upd_list)
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    48
next
61688
d04b1b4fb015 corrected inefficient implementation
nipkow
parents: 61686
diff changeset
    49
  case 3 thus ?case
d04b1b4fb015 corrected inefficient implementation
nipkow
parents: 61686
diff changeset
    50
    by(simp add: fun_eq_iff delete inv_delete map_of_del_list lookup sorted_del_list)
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    51
next
61686
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    52
  case 4 thus ?case by(simp add: empty inv_empty)
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    53
next
61688
d04b1b4fb015 corrected inefficient implementation
nipkow
parents: 61686
diff changeset
    54
  case 5 thus ?case by(simp add: update inv_update sorted_upd_list)
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    55
next
61686
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    56
  case 6 thus ?case by (auto simp: delete inv_delete sorted_del_list)
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    57
qed
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    58
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    59
end
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    60
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    61
end