src/HOL/Library/List_Cset.thy
author bulwahn
Fri, 20 Jan 2012 09:28:50 +0100
changeset 46305 8ea02e499d53
parent 46147 2c4d8de91c4c
child 47232 e2f0176149d0
permissions -rw-r--r--
adding check_all instance for sets; tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
     1
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
     2
(* Author: Florian Haftmann, TU Muenchen *)
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
     3
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
     4
header {* implementation of Cset.sets based on lists *}
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
     5
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
     6
theory List_Cset
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
     7
imports Cset
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
     8
begin
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
     9
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    10
code_datatype Cset.set Cset.coset
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    11
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    12
lemma member_code [code]:
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    13
  "member (Cset.set xs) = List.member xs"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    14
  "member (Cset.coset xs) = Not \<circ> List.member xs"
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    15
  by (simp_all add: fun_eq_iff List.member_def)
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    16
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    17
definition (in term_syntax)
46305
8ea02e499d53 adding check_all instance for sets; tuned
bulwahn
parents: 46147
diff changeset
    18
  csetify :: "'a\<Colon>typerep list \<times> (unit \<Rightarrow> Code_Evaluation.term)
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    19
    \<Rightarrow> 'a Cset.set \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
46305
8ea02e499d53 adding check_all instance for sets; tuned
bulwahn
parents: 46147
diff changeset
    20
  [code_unfold]: "csetify xs = Code_Evaluation.valtermify Cset.set {\<cdot>} xs"
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    21
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    22
notation fcomp (infixl "\<circ>>" 60)
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    23
notation scomp (infixl "\<circ>\<rightarrow>" 60)
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    24
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    25
instantiation Cset.set :: (random) random
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    26
begin
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    27
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    28
definition
46305
8ea02e499d53 adding check_all instance for sets; tuned
bulwahn
parents: 46147
diff changeset
    29
  "Quickcheck.random i = Quickcheck.random i \<circ>\<rightarrow> (\<lambda>xs. Pair (csetify xs))"
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    30
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    31
instance ..
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    32
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    33
end
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    34
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    35
no_notation fcomp (infixl "\<circ>>" 60)
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    36
no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    37
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    38
subsection {* Basic operations *}
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    39
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    40
lemma is_empty_set [code]:
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    41
  "Cset.is_empty (Cset.set xs) \<longleftrightarrow> List.null xs"
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    42
  by (simp add: is_empty_set null_def)
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    43
hide_fact (open) is_empty_set
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    44
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    45
lemma empty_set [code]:
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    46
  "Cset.empty = Cset.set []"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    47
  by simp
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    48
hide_fact (open) empty_set
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    49
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    50
lemma UNIV_set [code]:
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    51
  "top = Cset.coset []"
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    52
  by (simp add: Cset.coset_def)
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    53
hide_fact (open) UNIV_set
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    54
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    55
lemma remove_set [code]:
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    56
  "Cset.remove x (Cset.set xs) = Cset.set (removeAll x xs)"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    57
  "Cset.remove x (Cset.coset xs) = Cset.coset (List.insert x xs)"
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    58
  by (simp_all add: Cset.set_def Cset.coset_def Compl_insert)
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    59
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    60
lemma insert_set [code]:
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    61
  "Cset.insert x (Cset.set xs) = Cset.set (List.insert x xs)"
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    62
  "Cset.insert x (Cset.coset xs) = Cset.coset (removeAll x xs)"
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    63
  by (simp_all add: Cset.set_def Cset.coset_def)
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    64
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    65
declare compl_set [code]
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    66
declare compl_coset [code]
44560
1711be44e76a corrected slip
haftmann
parents: 44558
diff changeset
    67
declare subtract_remove [code]
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    68
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    69
lemma map_set [code]:
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    70
  "Cset.map f (Cset.set xs) = Cset.set (remdups (List.map f xs))"
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    71
  by (simp add: Cset.set_def)
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    72
  
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    73
lemma filter_set [code]:
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    74
  "Cset.filter P (Cset.set xs) = Cset.set (List.filter P xs)"
46147
2c4d8de91c4c moved lemmas about List.set and set operations to List theory
haftmann
parents: 45969
diff changeset
    75
  by (simp add: Cset.set_def)
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    76
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    77
lemma forall_set [code]:
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    78
  "Cset.forall P (Cset.set xs) \<longleftrightarrow> list_all P xs"
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    79
  by (simp add: Cset.set_def list_all_iff)
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    80
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    81
lemma exists_set [code]:
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    82
  "Cset.exists P (Cset.set xs) \<longleftrightarrow> list_ex P xs"
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    83
  by (simp add: Cset.set_def list_ex_iff)
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    84
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    85
lemma card_set [code]:
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    86
  "Cset.card (Cset.set xs) = length (remdups xs)"
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    87
proof -
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    88
  have "Finite_Set.card (set (remdups xs)) = length (remdups xs)"
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    89
    by (rule distinct_card) simp
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
    90
  then show ?thesis by (simp add: Cset.set_def)
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    91
qed
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    92
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    93
context complete_lattice
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    94
begin
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    95
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    96
declare Infimum_inf [code]
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
    97
declare Supremum_sup [code]
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    98
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
    99
end
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   100
43979
9f27d2bf4087 fixed code generator setup in List_Cset
Andreas Lochbihler
parents: 43971
diff changeset
   101
declare Cset.single_code [code del]
43971
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
   102
lemma single_set [code]:
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
   103
  "Cset.single a = Cset.set [a]"
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
   104
by(simp add: Cset.single_code)
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
   105
hide_fact (open) single_set
892030194015 added operations to Cset with code equations in backing implementations
Andreas Lochbihler
parents: 43307
diff changeset
   106
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
   107
declare Cset.bind_set [code]
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
   108
declare Cset.pred_of_cset_set [code]
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   109
44556
c0fd385a41f4 adapted to changes in Cset.thy
haftmann
parents: 43979
diff changeset
   110
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   111
subsection {* Derived operations *}
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   112
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   113
lemma subset_eq_forall [code]:
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   114
  "A \<le> B \<longleftrightarrow> Cset.forall (member B) A"
44556
c0fd385a41f4 adapted to changes in Cset.thy
haftmann
parents: 43979
diff changeset
   115
  by (simp add: subset_eq member_def)
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   116
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   117
lemma subset_subset_eq [code]:
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   118
  "A < B \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> (A :: 'a Cset.set)"
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   119
  by (fact less_le_not_le)
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   120
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   121
instantiation Cset.set :: (type) equal
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   122
begin
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   123
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   124
definition [code]:
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   125
  "HOL.equal A B \<longleftrightarrow> A \<le> B \<and> B \<le> (A :: 'a Cset.set)"
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   126
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   127
instance proof
45969
562e99c3d316 dropped references to obsolete fact `mem_def`
haftmann
parents: 44560
diff changeset
   128
qed (auto simp add: equal_set_def Cset.set_eq_iff Cset.member_def fun_eq_iff)
43241
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   129
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   130
end
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   131
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   132
lemma [code nbe]:
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   133
  "HOL.equal (A :: 'a Cset.set) A \<longleftrightarrow> True"
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   134
  by (fact equal_refl)
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   135
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   136
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   137
subsection {* Functorial operations *}
93b1183e43e5 splitting Cset into Cset and List_Cset
bulwahn
parents:
diff changeset
   138
44558
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
   139
declare inter_project [code]
cc878a312673 Cset, Dlist_Cset, List_Cset: restructured
haftmann
parents: 44556
diff changeset
   140
declare union_insert [code]
43307
1a32a953cef1 local simp rule in List_Cset
bulwahn
parents: 43241
diff changeset
   141
44556
c0fd385a41f4 adapted to changes in Cset.thy
haftmann
parents: 43979
diff changeset
   142
end