src/HOL/Library/Dlist_Cset.thy
author haftmann
Sat, 24 Dec 2011 15:53:10 +0100
changeset 45970 b6d0cff57d96
parent 44563 01b2732cf4ad
child 47232 e2f0176149d0
permissions -rw-r--r--
adjusted to set/pred distinction by means of type constructor `set`
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen *)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
     2
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
     3
header {* Canonical implementation of sets by distinct lists *}
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
     4
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
     5
theory Dlist_Cset
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
     6
imports Dlist Cset
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
     7
begin
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
     8
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
     9
definition Set :: "'a dlist \<Rightarrow> 'a Cset.set" where
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
    10
  "Set dxs = Cset.set (list_of_dlist dxs)"
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    11
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    12
definition Coset :: "'a dlist \<Rightarrow> 'a Cset.set" where
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    13
  "Coset dxs = Cset.coset (list_of_dlist dxs)"
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    14
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    15
code_datatype Set Coset
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    16
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    17
lemma Set_Dlist [simp]:
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    18
  "Set (Dlist xs) = Cset.set xs"
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    19
  by (rule Cset.set_eqI) (simp add: Set_def)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    20
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    21
lemma Coset_Dlist [simp]:
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    22
  "Coset (Dlist xs) = Cset.coset xs"
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    23
  by (rule Cset.set_eqI) (simp add: Coset_def)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    24
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    25
lemma member_Set [simp]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    26
  "Cset.member (Set dxs) = List.member (list_of_dlist dxs)"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    27
  by (simp add: Set_def fun_eq_iff List.member_def)
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    28
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    29
lemma member_Coset [simp]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    30
  "Cset.member (Coset dxs) = Not \<circ> List.member (list_of_dlist dxs)"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    31
  by (simp add: Coset_def fun_eq_iff List.member_def)
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    32
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    33
lemma Set_dlist_of_list [code]:
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
    34
  "Cset.set xs = Set (dlist_of_list xs)"
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    35
  by (rule Cset.set_eqI) simp
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    36
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    37
lemma Coset_dlist_of_list [code]:
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    38
  "Cset.coset xs = Coset (dlist_of_list xs)"
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    39
  by (rule Cset.set_eqI) simp
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    40
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    41
lemma is_empty_Set [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    42
  "Cset.is_empty (Set dxs) \<longleftrightarrow> Dlist.null dxs"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    43
  by (simp add: Dlist.null_def List.null_def Set_def)
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    44
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    45
lemma bot_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    46
  "bot = Set Dlist.empty"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    47
  by (simp add: empty_def)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    48
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    49
lemma top_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    50
  "top = Coset Dlist.empty"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    51
  by (simp add: empty_def Cset.coset_def)
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    52
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    53
lemma insert_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    54
  "Cset.insert x (Set dxs) = Set (Dlist.insert x dxs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    55
  "Cset.insert x (Coset dxs) = Coset (Dlist.remove x dxs)"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    56
  by (simp_all add: Dlist.insert_def Dlist.remove_def Cset.set_def Cset.coset_def Set_def Coset_def)
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    57
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    58
lemma remove_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    59
  "Cset.remove x (Set dxs) = Set (Dlist.remove x dxs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    60
  "Cset.remove x (Coset dxs) = Coset (Dlist.insert x dxs)"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    61
  by (simp_all add: Dlist.insert_def Dlist.remove_def Cset.set_def Cset.coset_def Set_def Coset_def Compl_insert)
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    62
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    63
lemma member_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    64
  "Cset.member (Set dxs) = Dlist.member dxs"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    65
  "Cset.member (Coset dxs) = Not \<circ> Dlist.member dxs"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    66
  by (simp_all add: List.member_def member_def fun_eq_iff Dlist.member_def)
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    67
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    68
lemma compl_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    69
  "- Set dxs = Coset dxs"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    70
  "- Coset dxs = Set dxs"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    71
  by (rule Cset.set_eqI, simp add: fun_eq_iff List.member_def Set_def Coset_def)+
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    72
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    73
lemma map_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    74
  "Cset.map f (Set dxs) = Set (Dlist.map f dxs)"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    75
  by (rule Cset.set_eqI) (simp add: fun_eq_iff List.member_def Set_def)
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    76
  
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    77
lemma filter_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    78
  "Cset.filter f (Set dxs) = Set (Dlist.filter f dxs)"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    79
  by (rule Cset.set_eqI) (simp add: fun_eq_iff List.member_def Set_def)
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    80
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    81
lemma forall_Set [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    82
  "Cset.forall P (Set xs) \<longleftrightarrow> list_all P (list_of_dlist xs)"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    83
  by (simp add: Set_def list_all_iff)
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    84
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    85
lemma exists_Set [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    86
  "Cset.exists P (Set xs) \<longleftrightarrow> list_ex P (list_of_dlist xs)"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    87
  by (simp add: Set_def list_ex_iff)
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    88
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    89
lemma card_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    90
  "Cset.card (Set dxs) = Dlist.length dxs"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
    91
  by (simp add: length_def Set_def distinct_card)
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    92
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    93
lemma inter_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    94
  "inf A (Set xs) = Set (Dlist.filter (Cset.member A) xs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    95
  "inf A (Coset xs) = Dlist.foldr Cset.remove xs A"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    96
  by (simp_all only: Set_def Coset_def foldr_def inter_project list_of_dlist_filter)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    97
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    98
lemma subtract_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    99
  "A - Set xs = Dlist.foldr Cset.remove xs A"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   100
  "A - Coset xs = Set (Dlist.filter (Cset.member A) xs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   101
  by (simp_all only: Set_def Coset_def foldr_def subtract_remove list_of_dlist_filter)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   102
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   103
lemma union_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   104
  "sup (Set xs) A = Dlist.foldr Cset.insert xs A"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   105
  "sup (Coset xs) A = Coset (Dlist.filter (Not \<circ> Cset.member A) xs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   106
  by (simp_all only: Set_def Coset_def foldr_def union_insert list_of_dlist_filter)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   107
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   108
context complete_lattice
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   109
begin
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   110
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   111
lemma Infimum_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   112
  "Infimum (Set As) = Dlist.foldr inf As top"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   113
  by (simp only: Set_def Infimum_inf foldr_def inf.commute)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   114
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   115
lemma Supremum_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   116
  "Supremum (Set As) = Dlist.foldr sup As bot"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   117
  by (simp only: Set_def Supremum_sup foldr_def sup.commute)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   118
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   119
end
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   120
44563
haftmann
parents: 44558
diff changeset
   121
declare Cset.single_code [code]
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   122
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   123
lemma bind_set [code]:
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
   124
  "Cset.bind (Dlist_Cset.Set xs) f = fold (sup \<circ> f) (list_of_dlist xs) Cset.empty"
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
   125
  by (simp add: Cset.bind_set Set_def)
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   126
hide_fact (open) bind_set
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   127
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   128
lemma pred_of_cset_set [code]:
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   129
  "pred_of_cset (Dlist_Cset.Set xs) = foldr sup (map Predicate.single (list_of_dlist xs)) bot"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 43971
diff changeset
   130
  by (simp add: Cset.pred_of_cset_set Dlist_Cset.Set_def)
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   131
hide_fact (open) pred_of_cset_set
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   132
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   133
end