src/HOL/Library/Dlist_Cset.thy
author huffman
Wed, 17 Aug 2011 15:12:34 -0700
changeset 44262 355d5438f5fb
parent 43971 892030194015
child 44558 cc878a312673
permissions -rw-r--r--
merged
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
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents: 43146
diff changeset
     6
imports Dlist List_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
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents: 43146
diff changeset
    13
  "Coset dxs = List_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
declare member_code [code del]
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents: 43146
diff changeset
    18
declare List_Cset.is_empty_set [code del]
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents: 43146
diff changeset
    19
declare List_Cset.empty_set [code del]
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents: 43146
diff changeset
    20
declare List_Cset.UNIV_set [code del]
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    21
declare insert_set [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    22
declare remove_set [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    23
declare compl_set [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    24
declare compl_coset [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    25
declare map_set [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    26
declare filter_set [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    27
declare forall_set [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    28
declare exists_set [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    29
declare card_set [code del]
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
    30
declare List_Cset.single_set [code del]
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
    31
declare List_Cset.bind_set [code del]
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
    32
declare List_Cset.pred_of_cset_set [code del]
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    33
declare inter_project [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    34
declare subtract_remove [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    35
declare union_insert [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    36
declare Infimum_inf [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    37
declare Supremum_sup [code del]
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    38
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    39
lemma Set_Dlist [simp]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    40
  "Set (Dlist xs) = Cset.Set (set xs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    41
  by (rule Cset.set_eqI) (simp add: Set_def)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    42
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    43
lemma Coset_Dlist [simp]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    44
  "Coset (Dlist xs) = Cset.Set (- set xs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    45
  by (rule Cset.set_eqI) (simp add: Coset_def)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    46
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    47
lemma member_Set [simp]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    48
  "Cset.member (Set dxs) = List.member (list_of_dlist dxs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    49
  by (simp add: Set_def member_set)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    50
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    51
lemma member_Coset [simp]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    52
  "Cset.member (Coset dxs) = Not \<circ> List.member (list_of_dlist dxs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    53
  by (simp add: Coset_def member_set not_set_compl)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    54
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    55
lemma Set_dlist_of_list [code]:
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
    56
  "Cset.set xs = Set (dlist_of_list xs)"
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    57
  by (rule Cset.set_eqI) simp
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    58
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    59
lemma Coset_dlist_of_list [code]:
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents: 43146
diff changeset
    60
  "List_Cset.coset xs = Coset (dlist_of_list xs)"
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    61
  by (rule Cset.set_eqI) simp
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 is_empty_Set [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    64
  "Cset.is_empty (Set dxs) \<longleftrightarrow> Dlist.null dxs"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    65
  by (simp add: Dlist.null_def List.null_def member_set)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    66
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    67
lemma bot_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    68
  "bot = Set Dlist.empty"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    69
  by (simp add: empty_def)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    70
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    71
lemma top_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    72
  "top = Coset Dlist.empty"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    73
  by (simp add: empty_def)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    74
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    75
lemma insert_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    76
  "Cset.insert x (Set dxs) = Set (Dlist.insert x dxs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    77
  "Cset.insert x (Coset dxs) = Coset (Dlist.remove x dxs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    78
  by (simp_all add: Dlist.insert_def Dlist.remove_def member_set not_set_compl)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    79
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    80
lemma remove_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    81
  "Cset.remove x (Set dxs) = Set (Dlist.remove x dxs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    82
  "Cset.remove x (Coset dxs) = Coset (Dlist.insert x dxs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    83
  by (auto simp add: Dlist.insert_def Dlist.remove_def member_set not_set_compl)
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 member_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    86
  "Cset.member (Set dxs) = Dlist.member dxs"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    87
  "Cset.member (Coset dxs) = Not \<circ> Dlist.member dxs"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    88
  by (simp_all add: member_def)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    89
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    90
lemma compl_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    91
  "- Set dxs = Coset dxs"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    92
  "- Coset dxs = Set dxs"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    93
  by (rule Cset.set_eqI, simp add: member_set not_set_compl)+
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    94
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    95
lemma map_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    96
  "Cset.map f (Set dxs) = Set (Dlist.map f dxs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    97
  by (rule Cset.set_eqI) (simp add: member_set)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    98
  
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
    99
lemma filter_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   100
  "Cset.filter f (Set dxs) = Set (Dlist.filter f dxs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   101
  by (rule Cset.set_eqI) (simp add: member_set)
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 forall_Set [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   104
  "Cset.forall P (Set xs) \<longleftrightarrow> list_all P (list_of_dlist xs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   105
  by (simp add: member_set list_all_iff)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   106
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   107
lemma exists_Set [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   108
  "Cset.exists P (Set xs) \<longleftrightarrow> list_ex P (list_of_dlist xs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   109
  by (simp add: member_set list_ex_iff)
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 card_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   112
  "Cset.card (Set dxs) = Dlist.length dxs"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   113
  by (simp add: length_def member_set distinct_card)
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 inter_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   116
  "inf A (Set xs) = Set (Dlist.filter (Cset.member A) xs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   117
  "inf A (Coset xs) = Dlist.foldr Cset.remove xs A"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   118
  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
   119
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   120
lemma subtract_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   121
  "A - Set xs = Dlist.foldr Cset.remove xs A"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   122
  "A - Coset xs = Set (Dlist.filter (Cset.member A) xs)"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   123
  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
   124
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   125
lemma union_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   126
  "sup (Set xs) A = Dlist.foldr Cset.insert xs A"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   127
  "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
   128
  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
   129
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   130
context complete_lattice
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   131
begin
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   132
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   133
lemma Infimum_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   134
  "Infimum (Set As) = Dlist.foldr inf As top"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   135
  by (simp only: Set_def Infimum_inf foldr_def inf.commute)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   136
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   137
lemma Supremum_code [code]:
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   138
  "Supremum (Set As) = Dlist.foldr sup As bot"
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   139
  by (simp only: Set_def Supremum_sup foldr_def sup.commute)
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   140
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   141
end
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   142
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   143
declare Cset.single_code[code]
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   144
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   145
lemma bind_set [code]:
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   146
  "Cset.bind (Dlist_Cset.Set xs) f = foldl (\<lambda>A x. sup A (f x)) Cset.empty (list_of_dlist xs)"
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   147
by(simp add: List_Cset.bind_set Dlist_Cset.Set_def)
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   148
hide_fact (open) bind_set
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   149
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   150
lemma pred_of_cset_set [code]:
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   151
  "pred_of_cset (Dlist_Cset.Set xs) = foldr sup (map Predicate.single (list_of_dlist xs)) bot"
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   152
by(simp add: List_Cset.pred_of_cset_set Dlist_Cset.Set_def)
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   153
hide_fact (open) pred_of_cset_set
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43241
diff changeset
   154
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents:
diff changeset
   155
end