| author | wenzelm |
| Wed, 15 Jun 2011 15:08:22 +0200 | |
| changeset 43395 | 85e468a8045a |
| parent 43241 | 93b1183e43e5 |
| child 43971 | 892030194015 |
| permissions | -rw-r--r-- |
| 31807 | 1 |
|
2 |
(* Author: Florian Haftmann, TU Muenchen *) |
|
3 |
||
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
4 |
header {* A dedicated set type which is executable on its finite part *}
|
| 31807 | 5 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
6 |
theory Cset |
|
37024
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
haftmann
parents:
37023
diff
changeset
|
7 |
imports More_Set More_List |
| 31807 | 8 |
begin |
9 |
||
10 |
subsection {* Lifting *}
|
|
11 |
||
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
12 |
typedef (open) 'a set = "UNIV :: 'a set set" |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
13 |
morphisms member Set by rule+ |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
14 |
hide_type (open) set |
| 31807 | 15 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
16 |
lemma member_Set [simp]: |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
17 |
"member (Set A) = A" |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
18 |
by (rule Set_inverse) rule |
| 34048 | 19 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
20 |
lemma Set_member [simp]: |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
21 |
"Set (member A) = A" |
| 37699 | 22 |
by (fact member_inverse) |
| 37468 | 23 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
24 |
lemma Set_inject [simp]: |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
25 |
"Set A = Set B \<longleftrightarrow> A = B" |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
26 |
by (simp add: Set_inject) |
| 37468 | 27 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
28 |
lemma set_eq_iff: |
|
39380
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
haftmann
parents:
39302
diff
changeset
|
29 |
"A = B \<longleftrightarrow> member A = member B" |
|
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
haftmann
parents:
39302
diff
changeset
|
30 |
by (simp add: member_inject) |
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
31 |
hide_fact (open) set_eq_iff |
|
39380
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
haftmann
parents:
39302
diff
changeset
|
32 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
33 |
lemma set_eqI: |
| 37473 | 34 |
"member A = member B \<Longrightarrow> A = B" |
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
35 |
by (simp add: Cset.set_eq_iff) |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
36 |
hide_fact (open) set_eqI |
| 37473 | 37 |
|
| 34048 | 38 |
subsection {* Lattice instantiation *}
|
39 |
||
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
40 |
instantiation Cset.set :: (type) boolean_algebra |
| 34048 | 41 |
begin |
42 |
||
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
43 |
definition less_eq_set :: "'a Cset.set \<Rightarrow> 'a Cset.set \<Rightarrow> bool" where |
| 34048 | 44 |
[simp]: "A \<le> B \<longleftrightarrow> member A \<subseteq> member B" |
45 |
||
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
46 |
definition less_set :: "'a Cset.set \<Rightarrow> 'a Cset.set \<Rightarrow> bool" where |
| 34048 | 47 |
[simp]: "A < B \<longleftrightarrow> member A \<subset> member B" |
48 |
||
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
49 |
definition inf_set :: "'a Cset.set \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
50 |
[simp]: "inf A B = Set (member A \<inter> member B)" |
| 34048 | 51 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
52 |
definition sup_set :: "'a Cset.set \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
53 |
[simp]: "sup A B = Set (member A \<union> member B)" |
| 34048 | 54 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
55 |
definition bot_set :: "'a Cset.set" where |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
56 |
[simp]: "bot = Set {}"
|
| 34048 | 57 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
58 |
definition top_set :: "'a Cset.set" where |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
59 |
[simp]: "top = Set UNIV" |
| 34048 | 60 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
61 |
definition uminus_set :: "'a Cset.set \<Rightarrow> 'a Cset.set" where |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
62 |
[simp]: "- A = Set (- (member A))" |
| 34048 | 63 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
64 |
definition minus_set :: "'a Cset.set \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
65 |
[simp]: "A - B = Set (member A - member B)" |
| 34048 | 66 |
|
67 |
instance proof |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
68 |
qed (auto intro: Cset.set_eqI) |
| 34048 | 69 |
|
70 |
end |
|
71 |
||
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
72 |
instantiation Cset.set :: (type) complete_lattice |
| 34048 | 73 |
begin |
74 |
||
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
75 |
definition Inf_set :: "'a Cset.set set \<Rightarrow> 'a Cset.set" where |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
76 |
[simp]: "Inf_set As = Set (Inf (image member As))" |
| 34048 | 77 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
78 |
definition Sup_set :: "'a Cset.set set \<Rightarrow> 'a Cset.set" where |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
79 |
[simp]: "Sup_set As = Set (Sup (image member As))" |
| 34048 | 80 |
|
81 |
instance proof |
|
82 |
qed (auto simp add: le_fun_def le_bool_def) |
|
83 |
||
84 |
end |
|
85 |
||
| 37023 | 86 |
|
| 31807 | 87 |
subsection {* Basic operations *}
|
88 |
||
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
89 |
definition is_empty :: "'a Cset.set \<Rightarrow> bool" where |
|
37024
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
haftmann
parents:
37023
diff
changeset
|
90 |
[simp]: "is_empty A \<longleftrightarrow> More_Set.is_empty (member A)" |
| 31807 | 91 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
92 |
definition insert :: "'a \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
93 |
[simp]: "insert x A = Set (Set.insert x (member A))" |
| 31807 | 94 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
95 |
definition remove :: "'a \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where |
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
96 |
[simp]: "remove x A = Set (More_Set.remove x (member A))" |
| 31807 | 97 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
98 |
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a Cset.set \<Rightarrow> 'b Cset.set" where
|
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
99 |
[simp]: "map f A = Set (image f (member A))" |
| 31807 | 100 |
|
|
41505
6d19301074cf
"enriched_type" replaces less specific "type_lifting"
haftmann
parents:
41372
diff
changeset
|
101 |
enriched_type map: map |
| 41372 | 102 |
by (simp_all add: fun_eq_iff image_compose) |
| 40604 | 103 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
104 |
definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where
|
|
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
105 |
[simp]: "filter P A = Set (More_Set.project P (member A))" |
| 31807 | 106 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
107 |
definition forall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Cset.set \<Rightarrow> bool" where
|
| 31846 | 108 |
[simp]: "forall P A \<longleftrightarrow> Ball (member A) P" |
| 31807 | 109 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
110 |
definition exists :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Cset.set \<Rightarrow> bool" where
|
| 31846 | 111 |
[simp]: "exists P A \<longleftrightarrow> Bex (member A) P" |
| 31807 | 112 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
113 |
definition card :: "'a Cset.set \<Rightarrow> nat" where |
| 31849 | 114 |
[simp]: "card A = Finite_Set.card (member A)" |
| 43241 | 115 |
|
| 34048 | 116 |
context complete_lattice |
117 |
begin |
|
| 31807 | 118 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
119 |
definition Infimum :: "'a Cset.set \<Rightarrow> 'a" where |
| 34048 | 120 |
[simp]: "Infimum A = Inf (member A)" |
| 31807 | 121 |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
122 |
definition Supremum :: "'a Cset.set \<Rightarrow> 'a" where |
| 34048 | 123 |
[simp]: "Supremum A = Sup (member A)" |
124 |
||
125 |
end |
|
| 31807 | 126 |
|
127 |
||
| 31846 | 128 |
subsection {* Simplified simprules *}
|
129 |
||
130 |
lemma is_empty_simp [simp]: |
|
131 |
"is_empty A \<longleftrightarrow> member A = {}"
|
|
|
37024
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
haftmann
parents:
37023
diff
changeset
|
132 |
by (simp add: More_Set.is_empty_def) |
| 31846 | 133 |
declare is_empty_def [simp del] |
134 |
||
135 |
lemma remove_simp [simp]: |
|
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
136 |
"remove x A = Set (member A - {x})"
|
|
37024
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
haftmann
parents:
37023
diff
changeset
|
137 |
by (simp add: More_Set.remove_def) |
| 31846 | 138 |
declare remove_def [simp del] |
139 |
||
| 31847 | 140 |
lemma filter_simp [simp]: |
|
40672
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents:
40604
diff
changeset
|
141 |
"filter P A = Set {x \<in> member A. P x}"
|
|
37024
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
haftmann
parents:
37023
diff
changeset
|
142 |
by (simp add: More_Set.project_def) |
| 31847 | 143 |
declare filter_def [simp del] |
| 31846 | 144 |
|
145 |
declare mem_def [simp del] |
|
146 |
||
| 31849 | 147 |
|
| 43241 | 148 |
hide_const (open) is_empty insert remove map filter forall exists card |
| 34048 | 149 |
Inter Union |
| 31849 | 150 |
|
| 31807 | 151 |
end |