src/HOL/Data_Structures/Map_Specs.thy
author desharna
Mon, 10 Oct 2022 19:07:54 +0200
changeset 76256 207b6fcfc47d
parent 68492 c7e0a7bcacaf
child 77224 e3e326a2dab5
permissions -rw-r--r--
removed unused universal variable from lemma reflp_onI
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
67965
aaa31cd0caef more name tuning
nipkow
parents: 67406
diff changeset
     3
section \<open>Specifications of Map ADT\<close>
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
     4
67965
aaa31cd0caef more name tuning
nipkow
parents: 67406
diff changeset
     5
theory Map_Specs
61640
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
67965
aaa31cd0caef more name tuning
nipkow
parents: 67406
diff changeset
     9
text \<open>The basic map interface with traditional \<open>set\<close>-based specification:\<close>
aaa31cd0caef more name tuning
nipkow
parents: 67406
diff changeset
    10
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    11
locale Map =
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    12
fixes empty :: "'m"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    13
fixes update :: "'a \<Rightarrow> 'b \<Rightarrow> 'm \<Rightarrow> 'm"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    14
fixes delete :: "'a \<Rightarrow> 'm \<Rightarrow> 'm"
62196
66fb3d1767f2 renamed map_of to lookup
nipkow
parents: 61688
diff changeset
    15
fixes lookup :: "'m \<Rightarrow> 'a \<Rightarrow> 'b option"
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    16
fixes invar :: "'m \<Rightarrow> bool"
62196
66fb3d1767f2 renamed map_of to lookup
nipkow
parents: 61688
diff changeset
    17
assumes map_empty: "lookup empty = (\<lambda>_. None)"
66fb3d1767f2 renamed map_of to lookup
nipkow
parents: 61688
diff changeset
    18
and map_update: "invar m \<Longrightarrow> lookup(update a b m) = (lookup m)(a := Some b)"
66fb3d1767f2 renamed map_of to lookup
nipkow
parents: 61688
diff changeset
    19
and map_delete: "invar m \<Longrightarrow> lookup(delete a m) = (lookup m)(a := None)"
61686
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    20
and invar_empty: "invar empty"
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    21
and invar_update: "invar m \<Longrightarrow> invar(update a b m)"
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    22
and invar_delete: "invar m \<Longrightarrow> invar(delete a m)"
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    23
68492
c7e0a7bcacaf added lemmas; uniform names
nipkow
parents: 68431
diff changeset
    24
lemmas (in Map) map_specs =
c7e0a7bcacaf added lemmas; uniform names
nipkow
parents: 68431
diff changeset
    25
  map_empty map_update map_delete invar_empty invar_update invar_delete
c7e0a7bcacaf added lemmas; uniform names
nipkow
parents: 68431
diff changeset
    26
67965
aaa31cd0caef more name tuning
nipkow
parents: 67406
diff changeset
    27
aaa31cd0caef more name tuning
nipkow
parents: 67406
diff changeset
    28
text \<open>The basic map interface with \<open>inorder\<close>-based specification:\<close>
aaa31cd0caef more name tuning
nipkow
parents: 67406
diff changeset
    29
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    30
locale Map_by_Ordered =
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    31
fixes empty :: "'t"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    32
fixes update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> 't \<Rightarrow> 't"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    33
fixes delete :: "'a \<Rightarrow> 't \<Rightarrow> 't"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    34
fixes lookup :: "'t \<Rightarrow> 'a \<Rightarrow> 'b option"
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    35
fixes inorder :: "'t \<Rightarrow> ('a * 'b) list"
61686
e6784939d645 tuned names
nipkow
parents: 61640
diff changeset
    36
fixes inv :: "'t \<Rightarrow> bool"
68431
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    37
assumes inorder_empty: "inorder empty = []"
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    38
and inorder_lookup: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    39
  lookup t a = map_of (inorder t) a"
68431
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    40
and inorder_update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    41
  inorder(update a b t) = upd_list a b (inorder t)"
68431
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    42
and inorder_delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow>
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    43
  inorder(delete a t) = del_list a (inorder t)"
68431
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    44
and inorder_inv_empty:  "inv empty"
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    45
and inorder_inv_update: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(update a b t)"
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    46
and inorder_inv_delete: "inv t \<and> sorted1 (inorder t) \<Longrightarrow> inv(delete a t)"
68492
c7e0a7bcacaf added lemmas; uniform names
nipkow
parents: 68431
diff changeset
    47
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    48
begin
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    49
67965
aaa31cd0caef more name tuning
nipkow
parents: 67406
diff changeset
    50
text \<open>It implements the traditional specification:\<close>
aaa31cd0caef more name tuning
nipkow
parents: 67406
diff changeset
    51
68431
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    52
definition invar :: "'t \<Rightarrow> bool" where
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    53
"invar t == inv t \<and> sorted1 (inorder t)"
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    54
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    55
sublocale Map
68431
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    56
  empty update delete lookup invar
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    57
proof(standard, goal_cases)
68431
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    58
  case 1 show ?case by (auto simp: inorder_lookup inorder_empty inorder_inv_empty)
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    59
next
61688
d04b1b4fb015 corrected inefficient implementation
nipkow
parents: 61686
diff changeset
    60
  case 2 thus ?case
68431
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    61
    by(simp add: fun_eq_iff inorder_update inorder_inv_update map_of_ins_list inorder_lookup
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    62
         sorted_upd_list invar_def)
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    63
next
61688
d04b1b4fb015 corrected inefficient implementation
nipkow
parents: 61686
diff changeset
    64
  case 3 thus ?case
68431
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    65
    by(simp add: fun_eq_iff inorder_delete inorder_inv_delete map_of_del_list inorder_lookup
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    66
         sorted_del_list invar_def)
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    67
next
68431
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    68
  case 4 thus ?case by(simp add: inorder_empty inorder_inv_empty invar_def)
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    69
next
68431
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    70
  case 5 thus ?case by(simp add: inorder_update inorder_inv_update sorted_upd_list invar_def)
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    71
next
b294e095f64c more abstract naming
nipkow
parents: 67965
diff changeset
    72
  case 6 thus ?case by (auto simp: inorder_delete inorder_inv_delete sorted_del_list invar_def)
61640
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    73
qed
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    74
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    75
end
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    76
44c9198f210c no CRLF
nipkow
parents: 61428
diff changeset
    77
end