src/HOL/Data_Structures/Set_by_Ordered.thy
author wenzelm
Wed, 04 Nov 2015 23:27:00 +0100
changeset 61578 6623c81cb15a
parent 61534 a88e07c8d0d5
child 61588 1d2907d0ed73
permissions -rw-r--r--
avoid ligatures;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     2
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     3
section {* Implementing Ordered Sets *}
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     4
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     5
theory Set_by_Ordered
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     6
imports List_Ins_Del
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     7
begin
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     8
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
     9
locale Set =
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    10
fixes empty :: "'s"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    11
fixes insert :: "'a \<Rightarrow> 's \<Rightarrow> 's"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    12
fixes delete :: "'a \<Rightarrow> 's \<Rightarrow> 's"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    13
fixes isin :: "'s \<Rightarrow> 'a \<Rightarrow> bool"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    14
fixes set :: "'s \<Rightarrow> 'a set"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    15
fixes invar :: "'s \<Rightarrow> bool"
61534
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    16
assumes set_empty:    "set empty = {}"
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    17
assumes set_isin:     "invar s \<Longrightarrow> isin s x = (x \<in> set s)"
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    18
assumes set_insert:   "invar s \<Longrightarrow> set(insert x s) = Set.insert x (set s)"
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    19
assumes set_delete:   "invar s \<Longrightarrow> set(delete x s) = set s - {x}"
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    20
assumes invar_empty:  "invar empty"
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    21
assumes invar_insert: "invar s \<Longrightarrow> invar(insert x s)"
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    22
assumes invar_delete: "invar s \<Longrightarrow> invar(delete x s)"
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    23
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    24
locale Set_by_Ordered =
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    25
fixes empty :: "'t"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    26
fixes insert :: "'a::linorder \<Rightarrow> 't \<Rightarrow> 't"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    27
fixes delete :: "'a \<Rightarrow> 't \<Rightarrow> 't"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    28
fixes isin :: "'t \<Rightarrow> 'a \<Rightarrow> bool"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    29
fixes inorder :: "'t \<Rightarrow> 'a list"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    30
fixes wf :: "'t \<Rightarrow> bool"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    31
assumes empty: "inorder empty = []"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    32
assumes isin: "wf t \<and> sorted(inorder t) \<Longrightarrow>
61534
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    33
  isin t x = (x \<in> elems (inorder t))"
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    34
assumes insert: "wf t \<and> sorted(inorder t) \<Longrightarrow>
61534
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    35
  inorder(insert x t) = ins_list x (inorder t)"
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    36
assumes delete: "wf t \<and> sorted(inorder t) \<Longrightarrow>
61534
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    37
  inorder(delete x t) = del_list x (inorder t)"
61428
5e1938107371 added invar empty
nipkow
parents: 61203
diff changeset
    38
assumes wf_empty:  "wf empty"
61534
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    39
assumes wf_insert: "wf t \<and> sorted(inorder t) \<Longrightarrow> wf(insert x t)"
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    40
assumes wf_delete: "wf t \<and> sorted(inorder t) \<Longrightarrow> wf(delete x t)"
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    41
begin
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    42
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    43
sublocale Set
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    44
  empty insert delete isin "elems o inorder" "\<lambda>t. wf t \<and> sorted(inorder t)"
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    45
proof(standard, goal_cases)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    46
  case 1 show ?case by (auto simp: empty)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    47
next
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    48
  case 2 thus ?case by(simp add: isin)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    49
next
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    50
  case 3 thus ?case by(simp add: insert)
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    51
next
61534
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    52
  case (4 s x) show ?case
a88e07c8d0d5 tuned names and optimized comparison order
nipkow
parents: 61428
diff changeset
    53
    using delete[OF 4, of x] 4 by (auto simp: distinct_if_sorted)
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    54
next
61428
5e1938107371 added invar empty
nipkow
parents: 61203
diff changeset
    55
  case 5 thus ?case by(simp add: empty wf_empty)
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    56
next
61428
5e1938107371 added invar empty
nipkow
parents: 61203
diff changeset
    57
  case 6 thus ?case by(simp add: insert wf_insert sorted_ins_list)
5e1938107371 added invar empty
nipkow
parents: 61203
diff changeset
    58
next
5e1938107371 added invar empty
nipkow
parents: 61203
diff changeset
    59
  case 7 thus ?case by (auto simp: delete wf_delete sorted_del_list)
61203
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    60
qed
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    61
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    62
end
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    63
a8a8eca85801 New subdirectory for functional data structures
nipkow
parents:
diff changeset
    64
end