src/HOL/Data_Structures/Map_by_Ordered.thy
author nipkow
Wed, 11 Nov 2015 18:32:26 +0100
changeset 61640 44c9198f210c
parent 61428 5e1938107371
child 61686 e6784939d645
permissions -rw-r--r--
no CRLF
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"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    15
assumes "map_of empty = (\<lambda>_. None)"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    16
assumes "invar m \<Longrightarrow> map_of(update a b m) = (map_of m)(a := Some b)"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    17
assumes "invar m \<Longrightarrow> map_of(delete a m) = (map_of m)(a := None)"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    18
assumes "invar empty"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    19
assumes "invar m \<Longrightarrow> invar(update a b m)"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    20
assumes "invar m \<Longrightarrow> invar(delete a m)"
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"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    28
fixes wf :: "'t \<Rightarrow> bool"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    29
assumes empty: "inorder empty = []"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    30
assumes lookup: "wf t \<and> sorted1 (inorder t) \<Longrightarrow>
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    31
  lookup t a = map_of (inorder t) a"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    32
assumes update: "wf t \<and> sorted1 (inorder t) \<Longrightarrow>
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    33
  inorder(update a b t) = upd_list a b (inorder t)"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    34
assumes delete: "wf t \<and> sorted1 (inorder t) \<Longrightarrow>
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    35
  inorder(delete a t) = del_list a (inorder t)"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    36
assumes wf_empty:  "wf empty"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    37
assumes wf_insert: "wf t \<and> sorted1 (inorder t) \<Longrightarrow> wf(update a b t)"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    38
assumes wf_delete: "wf t \<and> sorted1 (inorder t) \<Longrightarrow> wf(delete a t)"
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
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    42
  empty update delete "map_of o inorder" "\<lambda>t. wf t \<and> sorted1 (inorder t)"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    43
proof(standard, goal_cases)
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    44
  case 1 show ?case by (auto simp: empty)
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    45
next
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    46
  case 2 thus ?case by(simp add: update map_of_ins_list)
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    47
next
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    48
  case 3 thus ?case by(simp add: delete map_of_del_list)
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    49
next
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    50
  case 4 thus ?case by(simp add: empty wf_empty)
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    51
next
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    52
  case 5 thus ?case by(simp add: update wf_insert sorted_upd_list)
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    53
next
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    54
  case 6 thus ?case by (auto simp: delete wf_delete sorted_del_list)
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    55
qed
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    56
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    57
end
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    58
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    59
end