| author | blanchet | 
| Mon, 22 Aug 2011 15:02:45 +0200 | |
| changeset 44400 | 01b8b6fcd857 | 
| parent 43971 | 892030194015 | 
| child 44555 | da75ffe3d988 | 
| 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: 
40604diff
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: 
40604diff
changeset | 6 | theory Cset | 
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
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: 
40604diff
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: 
40604diff
changeset | 13 | morphisms member Set by rule+ | 
| 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 14 | hide_type (open) set | 
| 31807 | 15 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 16 | lemma member_Set [simp]: | 
| 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 17 | "member (Set A) = A" | 
| 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
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: 
40604diff
changeset | 20 | lemma Set_member [simp]: | 
| 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
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: 
40604diff
changeset | 24 | lemma Set_inject [simp]: | 
| 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 25 | "Set A = Set B \<longleftrightarrow> A = B" | 
| 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
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: 
40604diff
changeset | 28 | lemma set_eq_iff: | 
| 39380 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
39302diff
changeset | 29 | "A = B \<longleftrightarrow> member A = member B" | 
| 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
39302diff
changeset | 30 | by (simp add: member_inject) | 
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 31 | hide_fact (open) set_eq_iff | 
| 39380 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
39302diff
changeset | 32 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
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: 
40604diff
changeset | 35 | by (simp add: Cset.set_eq_iff) | 
| 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
changeset | 55 | definition bot_set :: "'a Cset.set" where | 
| 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 56 |   [simp]: "bot = Set {}"
 | 
| 34048 | 57 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 58 | definition top_set :: "'a Cset.set" where | 
| 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 59 | [simp]: "top = Set UNIV" | 
| 34048 | 60 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
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: 
40604diff
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 | ||
| 43971 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 89 | abbreviation empty :: "'a Cset.set" where "empty \<equiv> bot" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 90 | hide_const (open) empty | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 91 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 92 | abbreviation UNIV :: "'a Cset.set" where "UNIV \<equiv> top" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 93 | hide_const (open) UNIV | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 94 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 95 | definition is_empty :: "'a Cset.set \<Rightarrow> bool" where | 
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 96 | [simp]: "is_empty A \<longleftrightarrow> More_Set.is_empty (member A)" | 
| 31807 | 97 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 98 | 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: 
40604diff
changeset | 99 | [simp]: "insert x A = Set (Set.insert x (member A))" | 
| 31807 | 100 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 101 | 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: 
40604diff
changeset | 102 | [simp]: "remove x A = Set (More_Set.remove x (member A))" | 
| 31807 | 103 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 104 | 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: 
40604diff
changeset | 105 | [simp]: "map f A = Set (image f (member A))" | 
| 31807 | 106 | |
| 41505 
6d19301074cf
"enriched_type" replaces less specific "type_lifting"
 haftmann parents: 
41372diff
changeset | 107 | enriched_type map: map | 
| 41372 | 108 | by (simp_all add: fun_eq_iff image_compose) | 
| 40604 | 109 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 110 | 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: 
40604diff
changeset | 111 | [simp]: "filter P A = Set (More_Set.project P (member A))" | 
| 31807 | 112 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 113 | definition forall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Cset.set \<Rightarrow> bool" where
 | 
| 31846 | 114 | [simp]: "forall P A \<longleftrightarrow> Ball (member A) P" | 
| 31807 | 115 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 116 | definition exists :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Cset.set \<Rightarrow> bool" where
 | 
| 31846 | 117 | [simp]: "exists P A \<longleftrightarrow> Bex (member A) P" | 
| 31807 | 118 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 119 | definition card :: "'a Cset.set \<Rightarrow> nat" where | 
| 31849 | 120 | [simp]: "card A = Finite_Set.card (member A)" | 
| 43241 | 121 | |
| 34048 | 122 | context complete_lattice | 
| 123 | begin | |
| 31807 | 124 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 125 | definition Infimum :: "'a Cset.set \<Rightarrow> 'a" where | 
| 34048 | 126 | [simp]: "Infimum A = Inf (member A)" | 
| 31807 | 127 | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 128 | definition Supremum :: "'a Cset.set \<Rightarrow> 'a" where | 
| 34048 | 129 | [simp]: "Supremum A = Sup (member A)" | 
| 130 | ||
| 131 | end | |
| 31807 | 132 | |
| 43971 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 133 | subsection {* More operations *}
 | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 134 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 135 | text {* conversion from @{typ "'a list"} *}
 | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 136 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 137 | definition set :: "'a list \<Rightarrow> 'a Cset.set" where | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 138 | "set xs = Set (List.set xs)" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 139 | hide_const (open) set | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 140 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 141 | text {* conversion from @{typ "'a Predicate.pred"} *}
 | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 142 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 143 | definition pred_of_cset :: "'a Cset.set \<Rightarrow> 'a Predicate.pred" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 144 | where [code del]: "pred_of_cset = Predicate.Pred \<circ> Cset.member" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 145 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 146 | definition of_pred :: "'a Predicate.pred \<Rightarrow> 'a Cset.set" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 147 | where "of_pred = Cset.Set \<circ> Predicate.eval" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 148 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 149 | definition of_seq :: "'a Predicate.seq \<Rightarrow> 'a Cset.set" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 150 | where "of_seq = of_pred \<circ> Predicate.pred_of_seq" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 151 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 152 | text {* monad operations *}
 | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 153 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 154 | definition single :: "'a \<Rightarrow> 'a Cset.set" where | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 155 |   "single a = Set {a}"
 | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 156 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 157 | definition bind :: "'a Cset.set \<Rightarrow> ('a \<Rightarrow> 'b Cset.set) \<Rightarrow> 'b Cset.set"
 | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 158 | (infixl "\<guillemotright>=" 70) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 159 | where "A \<guillemotright>= f = Set (\<Union>x \<in> member A. member (f x))" | 
| 31807 | 160 | |
| 31846 | 161 | subsection {* Simplified simprules *}
 | 
| 162 | ||
| 43971 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 163 | lemma empty_simp [simp]: "member Cset.empty = {}"
 | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 164 | by(simp) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 165 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 166 | lemma UNIV_simp [simp]: "member Cset.UNIV = UNIV" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 167 | by simp | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 168 | |
| 31846 | 169 | lemma is_empty_simp [simp]: | 
| 170 |   "is_empty A \<longleftrightarrow> member A = {}"
 | |
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 171 | by (simp add: More_Set.is_empty_def) | 
| 31846 | 172 | declare is_empty_def [simp del] | 
| 173 | ||
| 174 | lemma remove_simp [simp]: | |
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 175 |   "remove x A = Set (member A - {x})"
 | 
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 176 | by (simp add: More_Set.remove_def) | 
| 31846 | 177 | declare remove_def [simp del] | 
| 178 | ||
| 31847 | 179 | lemma filter_simp [simp]: | 
| 40672 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 haftmann parents: 
40604diff
changeset | 180 |   "filter P A = Set {x \<in> member A. P x}"
 | 
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 181 | by (simp add: More_Set.project_def) | 
| 31847 | 182 | declare filter_def [simp del] | 
| 31846 | 183 | |
| 43971 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 184 | lemma member_set [simp]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 185 | "member (Cset.set xs) = set xs" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 186 | by (simp add: set_def) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 187 | hide_fact (open) member_set set_def | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 188 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 189 | lemma set_simps [simp]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 190 | "Cset.set [] = Cset.empty" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 191 | "Cset.set (x # xs) = insert x (Cset.set xs)" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 192 | by(simp_all add: Cset.set_def) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 193 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 194 | lemma member_SUPR [simp]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 195 | "member (SUPR A f) = SUPR A (member \<circ> f)" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 196 | unfolding SUPR_def by simp | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 197 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 198 | lemma member_bind [simp]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 199 | "member (P \<guillemotright>= f) = member (SUPR (member P) f)" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 200 | by (simp add: bind_def Cset.set_eq_iff) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 201 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 202 | lemma member_single [simp]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 203 |   "member (single a) = {a}"
 | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 204 | by(simp add: single_def) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 205 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 206 | lemma single_sup_simps [simp]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 207 | shows single_sup: "sup (single a) A = insert a A" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 208 | and sup_single: "sup A (single a) = insert a A" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 209 | by(auto simp add: Cset.set_eq_iff) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 210 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 211 | lemma single_bind [simp]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 212 | "single a \<guillemotright>= B = B a" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 213 | by(simp add: bind_def single_def) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 214 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 215 | lemma bind_bind: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 216 | "(A \<guillemotright>= B) \<guillemotright>= C = A \<guillemotright>= (\<lambda>x. B x \<guillemotright>= C)" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 217 | by(simp add: bind_def) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 218 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 219 | lemma bind_single [simp]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 220 | "A \<guillemotright>= single = A" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 221 | by(simp add: bind_def single_def) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 222 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 223 | lemma bind_const: "A \<guillemotright>= (\<lambda>_. B) = (if Cset.is_empty A then Cset.empty else B)" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 224 | by(auto simp add: Cset.set_eq_iff) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 225 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 226 | lemma empty_bind [simp]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 227 | "Cset.empty \<guillemotright>= f = Cset.empty" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 228 | by(simp add: Cset.set_eq_iff) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 229 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 230 | lemma member_of_pred [simp]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 231 |   "member (of_pred P) = {x. Predicate.eval P x}"
 | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 232 | by(simp add: of_pred_def Collect_def) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 233 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 234 | lemma member_of_seq [simp]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 235 |   "member (of_seq xq) = {x. Predicate.member xq x}"
 | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 236 | by(simp add: of_seq_def eval_member) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 237 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 238 | lemma eval_pred_of_cset [simp]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 239 | "Predicate.eval (pred_of_cset A) = Cset.member A" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 240 | by(simp add: pred_of_cset_def) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 241 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 242 | subsection {* Default implementations *}
 | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 243 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 244 | lemma set_code [code]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 245 | "Cset.set = foldl (\<lambda>A x. insert x A) Cset.empty" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 246 | proof(rule ext, rule Cset.set_eqI) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 247 | fix xs | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 248 | show "member (Cset.set xs) = member (foldl (\<lambda>A x. insert x A) Cset.empty xs)" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 249 | by(induct xs rule: rev_induct)(simp_all) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 250 | qed | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 251 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 252 | lemma single_code [code]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 253 | "single a = insert a Cset.empty" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 254 | by(simp add: Cset.single_def) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 255 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 256 | lemma of_pred_code [code]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 257 | "of_pred (Predicate.Seq f) = (case f () of | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 258 | Predicate.Empty \<Rightarrow> Cset.empty | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 259 | | Predicate.Insert x P \<Rightarrow> Cset.insert x (of_pred P) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 260 | | Predicate.Join P xq \<Rightarrow> sup (of_pred P) (of_seq xq))" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 261 | by(auto split: seq.split | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 262 | simp add: Predicate.Seq_def of_pred_def eval_member Cset.set_eq_iff) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 263 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 264 | lemma of_seq_code [code]: | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 265 | "of_seq Predicate.Empty = Cset.empty" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 266 | "of_seq (Predicate.Insert x P) = Cset.insert x (of_pred P)" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 267 | "of_seq (Predicate.Join P xq) = sup (of_pred P) (of_seq xq)" | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 268 | by(auto simp add: of_seq_def of_pred_def Cset.set_eq_iff) | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 269 | |
| 31846 | 270 | declare mem_def [simp del] | 
| 271 | ||
| 43971 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 272 | no_notation bind (infixl "\<guillemotright>=" 70) | 
| 31849 | 273 | |
| 43241 | 274 | hide_const (open) is_empty insert remove map filter forall exists card | 
| 43971 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 275 | Inter Union bind single of_pred of_seq | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 276 | |
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 277 | hide_fact (open) set_def pred_of_cset_def of_pred_def of_seq_def single_def | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 278 | bind_def empty_simp UNIV_simp member_set set_simps member_SUPR member_bind | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 279 | member_single single_sup_simps single_sup sup_single single_bind | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 280 | bind_bind bind_single bind_const empty_bind member_of_pred member_of_seq | 
| 
892030194015
added operations to Cset with code equations in backing implementations
 Andreas Lochbihler parents: 
43241diff
changeset | 281 | eval_pred_of_cset set_code single_code of_pred_code of_seq_code | 
| 31849 | 282 | |
| 31807 | 283 | end |