src/HOL/Library/Dlist.thy
author haftmann
Thu, 20 May 2010 16:35:53 +0200
changeset 37014 f9681d9d1d56
parent 36980 1a4cc941171a
child 37021 d754fb55a20f
permissions -rw-r--r--
moved generic List operations to theory More_List
haftmann@35303
     1
(* Author: Florian Haftmann, TU Muenchen *)
haftmann@35303
     2
haftmann@35303
     3
header {* Lists with elements distinct as canonical example for datatype invariants *}
haftmann@35303
     4
haftmann@35303
     5
theory Dlist
haftmann@37014
     6
imports Main More_List Fset
haftmann@35303
     7
begin
haftmann@35303
     8
haftmann@35303
     9
section {* The type of distinct lists *}
haftmann@35303
    10
haftmann@35303
    11
typedef (open) 'a dlist = "{xs::'a list. distinct xs}"
haftmann@35303
    12
  morphisms list_of_dlist Abs_dlist
haftmann@35303
    13
proof
haftmann@35303
    14
  show "[] \<in> ?dlist" by simp
haftmann@35303
    15
qed
haftmann@35303
    16
haftmann@36274
    17
lemma dlist_ext:
haftmann@36274
    18
  assumes "list_of_dlist xs = list_of_dlist ys"
haftmann@36274
    19
  shows "xs = ys"
haftmann@36274
    20
  using assms by (simp add: list_of_dlist_inject)
haftmann@36274
    21
haftmann@36112
    22
haftmann@35303
    23
text {* Formal, totalized constructor for @{typ "'a dlist"}: *}
haftmann@35303
    24
haftmann@35303
    25
definition Dlist :: "'a list \<Rightarrow> 'a dlist" where
haftmann@35303
    26
  [code del]: "Dlist xs = Abs_dlist (remdups xs)"
haftmann@35303
    27
haftmann@35303
    28
lemma distinct_list_of_dlist [simp]:
haftmann@35303
    29
  "distinct (list_of_dlist dxs)"
haftmann@35303
    30
  using list_of_dlist [of dxs] by simp
haftmann@35303
    31
haftmann@35303
    32
lemma list_of_dlist_Dlist [simp]:
haftmann@35303
    33
  "list_of_dlist (Dlist xs) = remdups xs"
haftmann@35303
    34
  by (simp add: Dlist_def Abs_dlist_inverse)
haftmann@35303
    35
haftmann@36112
    36
lemma Dlist_list_of_dlist [simp, code abstype]:
haftmann@35303
    37
  "Dlist (list_of_dlist dxs) = dxs"
haftmann@35303
    38
  by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id)
haftmann@35303
    39
haftmann@35303
    40
haftmann@35303
    41
text {* Fundamental operations: *}
haftmann@35303
    42
haftmann@35303
    43
definition empty :: "'a dlist" where
haftmann@35303
    44
  "empty = Dlist []"
haftmann@35303
    45
haftmann@35303
    46
definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
haftmann@35303
    47
  "insert x dxs = Dlist (List.insert x (list_of_dlist dxs))"
haftmann@35303
    48
haftmann@35303
    49
definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
haftmann@35303
    50
  "remove x dxs = Dlist (remove1 x (list_of_dlist dxs))"
haftmann@35303
    51
haftmann@35303
    52
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" where
haftmann@35303
    53
  "map f dxs = Dlist (remdups (List.map f (list_of_dlist dxs)))"
haftmann@35303
    54
haftmann@35303
    55
definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
haftmann@35303
    56
  "filter P dxs = Dlist (List.filter P (list_of_dlist dxs))"
haftmann@35303
    57
haftmann@35303
    58
haftmann@35303
    59
text {* Derived operations: *}
haftmann@35303
    60
haftmann@35303
    61
definition null :: "'a dlist \<Rightarrow> bool" where
haftmann@35303
    62
  "null dxs = List.null (list_of_dlist dxs)"
haftmann@35303
    63
haftmann@35303
    64
definition member :: "'a dlist \<Rightarrow> 'a \<Rightarrow> bool" where
haftmann@35303
    65
  "member dxs = List.member (list_of_dlist dxs)"
haftmann@35303
    66
haftmann@35303
    67
definition length :: "'a dlist \<Rightarrow> nat" where
haftmann@35303
    68
  "length dxs = List.length (list_of_dlist dxs)"
haftmann@35303
    69
haftmann@35303
    70
definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
haftmann@37014
    71
  "fold f dxs = More_List.fold f (list_of_dlist dxs)"
haftmann@37014
    72
haftmann@37014
    73
definition foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
haftmann@37014
    74
  "foldr f dxs = List.foldr f (list_of_dlist dxs)"
haftmann@35303
    75
haftmann@35303
    76
haftmann@35303
    77
section {* Executable version obeying invariant *}
haftmann@35303
    78
haftmann@35303
    79
lemma list_of_dlist_empty [simp, code abstract]:
haftmann@35303
    80
  "list_of_dlist empty = []"
haftmann@35303
    81
  by (simp add: empty_def)
haftmann@35303
    82
haftmann@35303
    83
lemma list_of_dlist_insert [simp, code abstract]:
haftmann@35303
    84
  "list_of_dlist (insert x dxs) = List.insert x (list_of_dlist dxs)"
haftmann@35303
    85
  by (simp add: insert_def)
haftmann@35303
    86
haftmann@35303
    87
lemma list_of_dlist_remove [simp, code abstract]:
haftmann@35303
    88
  "list_of_dlist (remove x dxs) = remove1 x (list_of_dlist dxs)"
haftmann@35303
    89
  by (simp add: remove_def)
haftmann@35303
    90
haftmann@35303
    91
lemma list_of_dlist_map [simp, code abstract]:
haftmann@35303
    92
  "list_of_dlist (map f dxs) = remdups (List.map f (list_of_dlist dxs))"
haftmann@35303
    93
  by (simp add: map_def)
haftmann@35303
    94
haftmann@35303
    95
lemma list_of_dlist_filter [simp, code abstract]:
haftmann@35303
    96
  "list_of_dlist (filter P dxs) = List.filter P (list_of_dlist dxs)"
haftmann@35303
    97
  by (simp add: filter_def)
haftmann@35303
    98
haftmann@35303
    99
haftmann@36980
   100
text {* Explicit executable conversion *}
haftmann@36980
   101
haftmann@36980
   102
definition dlist_of_list [simp]:
haftmann@36980
   103
  "dlist_of_list = Dlist"
haftmann@36980
   104
haftmann@36980
   105
lemma [code abstract]:
haftmann@36980
   106
  "list_of_dlist (dlist_of_list xs) = remdups xs"
haftmann@36980
   107
  by simp
haftmann@36980
   108
haftmann@36980
   109
haftmann@35303
   110
section {* Implementation of sets by distinct lists -- canonical! *}
haftmann@35303
   111
haftmann@35303
   112
definition Set :: "'a dlist \<Rightarrow> 'a fset" where
haftmann@35303
   113
  "Set dxs = Fset.Set (list_of_dlist dxs)"
haftmann@35303
   114
haftmann@35303
   115
definition Coset :: "'a dlist \<Rightarrow> 'a fset" where
haftmann@35303
   116
  "Coset dxs = Fset.Coset (list_of_dlist dxs)"
haftmann@35303
   117
haftmann@35303
   118
code_datatype Set Coset
haftmann@35303
   119
haftmann@35303
   120
declare member_code [code del]
haftmann@35303
   121
declare is_empty_Set [code del]
haftmann@35303
   122
declare empty_Set [code del]
haftmann@35303
   123
declare UNIV_Set [code del]
haftmann@35303
   124
declare insert_Set [code del]
haftmann@35303
   125
declare remove_Set [code del]
haftmann@35303
   126
declare map_Set [code del]
haftmann@35303
   127
declare filter_Set [code del]
haftmann@35303
   128
declare forall_Set [code del]
haftmann@35303
   129
declare exists_Set [code del]
haftmann@35303
   130
declare card_Set [code del]
haftmann@35303
   131
declare inter_project [code del]
haftmann@35303
   132
declare subtract_remove [code del]
haftmann@35303
   133
declare union_insert [code del]
haftmann@35303
   134
declare Infimum_inf [code del]
haftmann@35303
   135
declare Supremum_sup [code del]
haftmann@35303
   136
haftmann@35303
   137
lemma Set_Dlist [simp]:
haftmann@35303
   138
  "Set (Dlist xs) = Fset (set xs)"
haftmann@35303
   139
  by (simp add: Set_def Fset.Set_def)
haftmann@35303
   140
haftmann@35303
   141
lemma Coset_Dlist [simp]:
haftmann@35303
   142
  "Coset (Dlist xs) = Fset (- set xs)"
haftmann@35303
   143
  by (simp add: Coset_def Fset.Coset_def)
haftmann@35303
   144
haftmann@35303
   145
lemma member_Set [simp]:
haftmann@35303
   146
  "Fset.member (Set dxs) = List.member (list_of_dlist dxs)"
haftmann@35303
   147
  by (simp add: Set_def member_set)
haftmann@35303
   148
haftmann@35303
   149
lemma member_Coset [simp]:
haftmann@35303
   150
  "Fset.member (Coset dxs) = Not \<circ> List.member (list_of_dlist dxs)"
haftmann@35303
   151
  by (simp add: Coset_def member_set not_set_compl)
haftmann@35303
   152
haftmann@36980
   153
lemma Set_dlist_of_list [code]:
haftmann@36980
   154
  "Fset.Set xs = Set (dlist_of_list xs)"
haftmann@36980
   155
  by simp
haftmann@36980
   156
haftmann@36980
   157
lemma Coset_dlist_of_list [code]:
haftmann@36980
   158
  "Fset.Coset xs = Coset (dlist_of_list xs)"
haftmann@36980
   159
  by simp
haftmann@36980
   160
haftmann@35303
   161
lemma is_empty_Set [code]:
haftmann@35303
   162
  "Fset.is_empty (Set dxs) \<longleftrightarrow> null dxs"
haftmann@35303
   163
  by (simp add: null_def null_empty member_set)
haftmann@35303
   164
haftmann@35303
   165
lemma bot_code [code]:
haftmann@35303
   166
  "bot = Set empty"
haftmann@35303
   167
  by (simp add: empty_def)
haftmann@35303
   168
haftmann@35303
   169
lemma top_code [code]:
haftmann@35303
   170
  "top = Coset empty"
haftmann@35303
   171
  by (simp add: empty_def)
haftmann@35303
   172
haftmann@35303
   173
lemma insert_code [code]:
haftmann@35303
   174
  "Fset.insert x (Set dxs) = Set (insert x dxs)"
haftmann@35303
   175
  "Fset.insert x (Coset dxs) = Coset (remove x dxs)"
haftmann@35303
   176
  by (simp_all add: insert_def remove_def member_set not_set_compl)
haftmann@35303
   177
haftmann@35303
   178
lemma remove_code [code]:
haftmann@35303
   179
  "Fset.remove x (Set dxs) = Set (remove x dxs)"
haftmann@35303
   180
  "Fset.remove x (Coset dxs) = Coset (insert x dxs)"
haftmann@35303
   181
  by (auto simp add: insert_def remove_def member_set not_set_compl)
haftmann@35303
   182
haftmann@35303
   183
lemma member_code [code]:
haftmann@35303
   184
  "Fset.member (Set dxs) = member dxs"
haftmann@35303
   185
  "Fset.member (Coset dxs) = Not \<circ> member dxs"
haftmann@35303
   186
  by (simp_all add: member_def)
haftmann@35303
   187
haftmann@35303
   188
lemma map_code [code]:
haftmann@35303
   189
  "Fset.map f (Set dxs) = Set (map f dxs)"
haftmann@35303
   190
  by (simp add: member_set)
haftmann@35303
   191
  
haftmann@35303
   192
lemma filter_code [code]:
haftmann@35303
   193
  "Fset.filter f (Set dxs) = Set (filter f dxs)"
haftmann@35303
   194
  by (simp add: member_set)
haftmann@35303
   195
haftmann@35303
   196
lemma forall_Set [code]:
haftmann@35303
   197
  "Fset.forall P (Set xs) \<longleftrightarrow> list_all P (list_of_dlist xs)"
haftmann@35303
   198
  by (simp add: member_set list_all_iff)
haftmann@35303
   199
haftmann@35303
   200
lemma exists_Set [code]:
haftmann@35303
   201
  "Fset.exists P (Set xs) \<longleftrightarrow> list_ex P (list_of_dlist xs)"
haftmann@35303
   202
  by (simp add: member_set list_ex_iff)
haftmann@35303
   203
haftmann@35303
   204
lemma card_code [code]:
haftmann@35303
   205
  "Fset.card (Set dxs) = length dxs"
haftmann@35303
   206
  by (simp add: length_def member_set distinct_card)
haftmann@35303
   207
haftmann@35303
   208
lemma inter_code [code]:
haftmann@35303
   209
  "inf A (Set xs) = Set (filter (Fset.member A) xs)"
haftmann@37014
   210
  "inf A (Coset xs) = foldr Fset.remove xs A"
haftmann@37014
   211
  by (simp_all only: Set_def Coset_def foldr_def inter_project list_of_dlist_filter)
haftmann@35303
   212
haftmann@35303
   213
lemma subtract_code [code]:
haftmann@37014
   214
  "A - Set xs = foldr Fset.remove xs A"
haftmann@35303
   215
  "A - Coset xs = Set (filter (Fset.member A) xs)"
haftmann@37014
   216
  by (simp_all only: Set_def Coset_def foldr_def subtract_remove list_of_dlist_filter)
haftmann@35303
   217
haftmann@35303
   218
lemma union_code [code]:
haftmann@37014
   219
  "sup (Set xs) A = foldr Fset.insert xs A"
haftmann@35303
   220
  "sup (Coset xs) A = Coset (filter (Not \<circ> Fset.member A) xs)"
haftmann@37014
   221
  by (simp_all only: Set_def Coset_def foldr_def union_insert list_of_dlist_filter)
haftmann@35303
   222
haftmann@35303
   223
context complete_lattice
haftmann@35303
   224
begin
haftmann@35303
   225
haftmann@35303
   226
lemma Infimum_code [code]:
haftmann@37014
   227
  "Infimum (Set As) = foldr inf As top"
haftmann@37014
   228
  by (simp only: Set_def Infimum_inf foldr_def inf.commute)
haftmann@35303
   229
haftmann@35303
   230
lemma Supremum_code [code]:
haftmann@37014
   231
  "Supremum (Set As) = foldr sup As bot"
haftmann@37014
   232
  by (simp only: Set_def Supremum_sup foldr_def sup.commute)
haftmann@35303
   233
haftmann@35303
   234
end
haftmann@35303
   235
haftmann@37014
   236
hide_const (open) member fold foldr empty insert remove map filter null member length fold
haftmann@35303
   237
haftmann@35303
   238
end