| author | wenzelm | 
| Thu, 11 Jul 2013 14:56:58 +0200 | |
| changeset 52597 | a8a81453833d | 
| parent 52435 | 6646bb548c6b | 
| child 53015 | a1119cf551e8 | 
| permissions | -rw-r--r-- | 
| 31596 | 1 | (* Author: Florian Haftmann, TU Muenchen *) | 
| 26348 | 2 | |
| 3 | header {* Finite types as explicit enumerations *}
 | |
| 4 | ||
| 5 | theory Enum | |
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49950diff
changeset | 6 | imports Map | 
| 26348 | 7 | begin | 
| 8 | ||
| 9 | subsection {* Class @{text enum} *}
 | |
| 10 | ||
| 29797 | 11 | class enum = | 
| 26348 | 12 | fixes enum :: "'a list" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 13 |   fixes enum_all :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 14 |   fixes enum_ex :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 33635 | 15 | assumes UNIV_enum: "UNIV = set enum" | 
| 26444 | 16 | and enum_distinct: "distinct enum" | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 17 | assumes enum_all_UNIV: "enum_all P \<longleftrightarrow> Ball UNIV P" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 18 | assumes enum_ex_UNIV: "enum_ex P \<longleftrightarrow> Bex UNIV P" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 19 |    -- {* tailored towards simple instantiation *}
 | 
| 26348 | 20 | begin | 
| 21 | ||
| 29797 | 22 | subclass finite proof | 
| 23 | qed (simp add: UNIV_enum) | |
| 26444 | 24 | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 25 | lemma enum_UNIV: | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 26 | "set enum = UNIV" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 27 | by (simp only: UNIV_enum) | 
| 26444 | 28 | |
| 40683 | 29 | lemma in_enum: "x \<in> set enum" | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 30 | by (simp add: enum_UNIV) | 
| 26348 | 31 | |
| 32 | lemma enum_eq_I: | |
| 33 | assumes "\<And>x. x \<in> set xs" | |
| 34 | shows "set enum = set xs" | |
| 35 | proof - | |
| 36 | from assms UNIV_eq_I have "UNIV = set xs" by auto | |
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 37 | with enum_UNIV show ?thesis by simp | 
| 26348 | 38 | qed | 
| 39 | ||
| 49972 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49950diff
changeset | 40 | lemma card_UNIV_length_enum: | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49950diff
changeset | 41 | "card (UNIV :: 'a set) = length enum" | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49950diff
changeset | 42 | by (simp add: UNIV_enum distinct_card enum_distinct) | 
| 
f11f8905d9fd
incorporated constant chars into instantiation proof for enum;
 haftmann parents: 
49950diff
changeset | 43 | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 44 | lemma enum_all [simp]: | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 45 | "enum_all = HOL.All" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 46 | by (simp add: fun_eq_iff enum_all_UNIV) | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 47 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 48 | lemma enum_ex [simp]: | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 49 | "enum_ex = HOL.Ex" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 50 | by (simp add: fun_eq_iff enum_ex_UNIV) | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 51 | |
| 26348 | 52 | end | 
| 53 | ||
| 54 | ||
| 49949 | 55 | subsection {* Implementations using @{class enum} *}
 | 
| 56 | ||
| 57 | subsubsection {* Unbounded operations and quantifiers *}
 | |
| 58 | ||
| 59 | lemma Collect_code [code]: | |
| 60 | "Collect P = set (filter P enum)" | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 61 | by (simp add: enum_UNIV) | 
| 49949 | 62 | |
| 50567 
768a3fbe4149
providing a custom code equation for vimage to overwrite the vimage definition that would be rewritten by set_comprehension_pointfree simproc in the code preprocessor to an non-terminating code equation
 bulwahn parents: 
49972diff
changeset | 63 | lemma vimage_code [code]: | 
| 
768a3fbe4149
providing a custom code equation for vimage to overwrite the vimage definition that would be rewritten by set_comprehension_pointfree simproc in the code preprocessor to an non-terminating code equation
 bulwahn parents: 
49972diff
changeset | 64 | "f -` B = set (filter (%x. f x : B) enum_class.enum)" | 
| 
768a3fbe4149
providing a custom code equation for vimage to overwrite the vimage definition that would be rewritten by set_comprehension_pointfree simproc in the code preprocessor to an non-terminating code equation
 bulwahn parents: 
49972diff
changeset | 65 | unfolding vimage_def Collect_code .. | 
| 
768a3fbe4149
providing a custom code equation for vimage to overwrite the vimage definition that would be rewritten by set_comprehension_pointfree simproc in the code preprocessor to an non-terminating code equation
 bulwahn parents: 
49972diff
changeset | 66 | |
| 49949 | 67 | definition card_UNIV :: "'a itself \<Rightarrow> nat" | 
| 68 | where | |
| 69 |   [code del]: "card_UNIV TYPE('a) = card (UNIV :: 'a set)"
 | |
| 70 | ||
| 71 | lemma [code]: | |
| 72 |   "card_UNIV TYPE('a :: enum) = card (set (Enum.enum :: 'a list))"
 | |
| 73 | by (simp only: card_UNIV_def enum_UNIV) | |
| 74 | ||
| 75 | lemma all_code [code]: "(\<forall>x. P x) \<longleftrightarrow> enum_all P" | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 76 | by simp | 
| 49949 | 77 | |
| 78 | lemma exists_code [code]: "(\<exists>x. P x) \<longleftrightarrow> enum_ex P" | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 79 | by simp | 
| 49949 | 80 | |
| 81 | lemma exists1_code [code]: "(\<exists>!x. P x) \<longleftrightarrow> list_ex1 P enum" | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 82 | by (auto simp add: list_ex1_iff enum_UNIV) | 
| 49949 | 83 | |
| 84 | ||
| 85 | subsubsection {* An executable choice operator *}
 | |
| 86 | ||
| 87 | definition | |
| 88 | [code del]: "enum_the = The" | |
| 89 | ||
| 90 | lemma [code]: | |
| 91 | "The P = (case filter P enum of [x] => x | _ => enum_the P)" | |
| 92 | proof - | |
| 93 |   {
 | |
| 94 | fix a | |
| 95 | assume filter_enum: "filter P enum = [a]" | |
| 96 | have "The P = a" | |
| 97 | proof (rule the_equality) | |
| 98 | fix x | |
| 99 | assume "P x" | |
| 100 | show "x = a" | |
| 101 | proof (rule ccontr) | |
| 102 | assume "x \<noteq> a" | |
| 103 | from filter_enum obtain us vs | |
| 104 | where enum_eq: "enum = us @ [a] @ vs" | |
| 105 | and "\<forall> x \<in> set us. \<not> P x" | |
| 106 | and "\<forall> x \<in> set vs. \<not> P x" | |
| 107 | and "P a" | |
| 108 | by (auto simp add: filter_eq_Cons_iff) (simp only: filter_empty_conv[symmetric]) | |
| 109 | with `P x` in_enum[of x, unfolded enum_eq] `x \<noteq> a` show "False" by auto | |
| 110 | qed | |
| 111 | next | |
| 112 | from filter_enum show "P a" by (auto simp add: filter_eq_Cons_iff) | |
| 113 | qed | |
| 114 | } | |
| 115 | from this show ?thesis | |
| 116 | unfolding enum_the_def by (auto split: list.split) | |
| 117 | qed | |
| 118 | ||
| 119 | code_abort enum_the | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
50567diff
changeset | 120 | |
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
50567diff
changeset | 121 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
50567diff
changeset | 122 | constant enum_the \<rightharpoonup> (Eval) "(fn '_ => raise Match)" | 
| 49949 | 123 | |
| 124 | ||
| 125 | subsubsection {* Equality and order on functions *}
 | |
| 26348 | 126 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 127 | instantiation "fun" :: (enum, equal) equal | 
| 26513 | 128 | begin | 
| 26348 | 129 | |
| 26513 | 130 | definition | 
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 131 | "HOL.equal f g \<longleftrightarrow> (\<forall>x \<in> set enum. f x = g x)" | 
| 26513 | 132 | |
| 31464 | 133 | instance proof | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 134 | qed (simp_all add: equal_fun_def fun_eq_iff enum_UNIV) | 
| 26513 | 135 | |
| 136 | end | |
| 26348 | 137 | |
| 40898 
882e860a1e83
changed order of lemmas to overwrite the general code equation with the nbe-specific one
 bulwahn parents: 
40683diff
changeset | 138 | lemma [code]: | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 139 | "HOL.equal f g \<longleftrightarrow> enum_all (%x. f x = g x)" | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 140 | by (auto simp add: equal fun_eq_iff) | 
| 40898 
882e860a1e83
changed order of lemmas to overwrite the general code equation with the nbe-specific one
 bulwahn parents: 
40683diff
changeset | 141 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 142 | lemma [code nbe]: | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 143 | "HOL.equal (f :: _ \<Rightarrow> _) f \<longleftrightarrow> True" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 144 | by (fact equal_refl) | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 145 | |
| 28562 | 146 | lemma order_fun [code]: | 
| 26348 | 147 | fixes f g :: "'a\<Colon>enum \<Rightarrow> 'b\<Colon>order" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 148 | shows "f \<le> g \<longleftrightarrow> enum_all (\<lambda>x. f x \<le> g x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 149 | and "f < g \<longleftrightarrow> f \<le> g \<and> enum_ex (\<lambda>x. f x \<noteq> g x)" | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 150 | by (simp_all add: fun_eq_iff le_fun_def order_less_le) | 
| 26968 | 151 | |
| 152 | ||
| 49949 | 153 | subsubsection {* Operations on relations *}
 | 
| 154 | ||
| 155 | lemma [code]: | |
| 156 | "Id = image (\<lambda>x. (x, x)) (set Enum.enum)" | |
| 157 | by (auto intro: imageI in_enum) | |
| 26968 | 158 | |
| 49949 | 159 | lemma tranclp_unfold [code, no_atp]: | 
| 160 |   "tranclp r a b \<longleftrightarrow> (a, b) \<in> trancl {(x, y). r x y}"
 | |
| 161 | by (simp add: trancl_def) | |
| 162 | ||
| 163 | lemma rtranclp_rtrancl_eq [code, no_atp]: | |
| 164 |   "rtranclp r x y \<longleftrightarrow> (x, y) \<in> rtrancl {(x, y). r x y}"
 | |
| 165 | by (simp add: rtrancl_def) | |
| 26968 | 166 | |
| 49949 | 167 | lemma max_ext_eq [code]: | 
| 168 |   "max_ext R = {(X, Y). finite X \<and> finite Y \<and> Y \<noteq> {} \<and> (\<forall>x. x \<in> X \<longrightarrow> (\<exists>xa \<in> Y. (x, xa) \<in> R))}"
 | |
| 169 | by (auto simp add: max_ext.simps) | |
| 170 | ||
| 171 | lemma max_extp_eq [code]: | |
| 172 |   "max_extp r x y \<longleftrightarrow> (x, y) \<in> max_ext {(x, y). r x y}"
 | |
| 173 | by (simp add: max_ext_def) | |
| 26348 | 174 | |
| 49949 | 175 | lemma mlex_eq [code]: | 
| 176 |   "f <*mlex*> R = {(x, y). f x < f y \<or> (f x \<le> f y \<and> (x, y) \<in> R)}"
 | |
| 177 | by (auto simp add: mlex_prod_def) | |
| 178 | ||
| 179 | lemma [code]: | |
| 180 |   fixes xs :: "('a::finite \<times> 'a) list"
 | |
| 181 |   shows "acc (set xs) = bacc (set xs) (card_UNIV TYPE('a))"
 | |
| 182 | by (simp add: card_UNIV_def acc_bacc_eq) | |
| 183 | ||
| 184 | lemma [code]: | |
| 185 |   "accp r = (\<lambda>x. x \<in> acc {(x, y). r x y})"
 | |
| 186 | by (simp add: acc_def) | |
| 40652 | 187 | |
| 26348 | 188 | |
| 49949 | 189 | subsection {* Default instances for @{class enum} *}
 | 
| 26348 | 190 | |
| 26444 | 191 | lemma map_of_zip_enum_is_Some: | 
| 192 | assumes "length ys = length (enum \<Colon> 'a\<Colon>enum list)" | |
| 193 | shows "\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y" | |
| 194 | proof - | |
| 195 | from assms have "x \<in> set (enum \<Colon> 'a\<Colon>enum list) \<longleftrightarrow> | |
| 196 | (\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y)" | |
| 197 | by (auto intro!: map_of_zip_is_Some) | |
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 198 | then show ?thesis using enum_UNIV by auto | 
| 26444 | 199 | qed | 
| 200 | ||
| 201 | lemma map_of_zip_enum_inject: | |
| 202 | fixes xs ys :: "'b\<Colon>enum list" | |
| 203 | assumes length: "length xs = length (enum \<Colon> 'a\<Colon>enum list)" | |
| 204 | "length ys = length (enum \<Colon> 'a\<Colon>enum list)" | |
| 205 | and map_of: "the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys)" | |
| 206 | shows "xs = ys" | |
| 207 | proof - | |
| 208 | have "map_of (zip (enum \<Colon> 'a list) xs) = map_of (zip (enum \<Colon> 'a list) ys)" | |
| 209 | proof | |
| 210 | fix x :: 'a | |
| 211 | from length map_of_zip_enum_is_Some obtain y1 y2 | |
| 212 | where "map_of (zip (enum \<Colon> 'a list) xs) x = Some y1" | |
| 213 | and "map_of (zip (enum \<Colon> 'a list) ys) x = Some y2" by blast | |
| 47230 | 214 | moreover from map_of | 
| 215 | have "the (map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) x) = the (map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x)" | |
| 26444 | 216 | by (auto dest: fun_cong) | 
| 217 | ultimately show "map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) x = map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x" | |
| 218 | by simp | |
| 219 | qed | |
| 220 | with length enum_distinct show "xs = ys" by (rule map_of_zip_inject) | |
| 221 | qed | |
| 222 | ||
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 223 | definition all_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
 | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 224 | where | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 225 | "all_n_lists P n \<longleftrightarrow> (\<forall>xs \<in> set (List.n_lists n enum). P xs)" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 226 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 227 | lemma [code]: | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 228 | "all_n_lists P n \<longleftrightarrow> (if n = 0 then P [] else enum_all (%x. all_n_lists (%xs. P (x # xs)) (n - 1)))" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 229 | unfolding all_n_lists_def enum_all | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 230 | by (cases n) (auto simp add: enum_UNIV) | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 231 | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 232 | definition ex_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
 | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 233 | where | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 234 | "ex_n_lists P n \<longleftrightarrow> (\<exists>xs \<in> set (List.n_lists n enum). P xs)" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 235 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 236 | lemma [code]: | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 237 | "ex_n_lists P n \<longleftrightarrow> (if n = 0 then P [] else enum_ex (%x. ex_n_lists (%xs. P (x # xs)) (n - 1)))" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 238 | unfolding ex_n_lists_def enum_ex | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 239 | by (cases n) (auto simp add: enum_UNIV) | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 240 | |
| 26444 | 241 | instantiation "fun" :: (enum, enum) enum | 
| 242 | begin | |
| 243 | ||
| 244 | definition | |
| 49948 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
48123diff
changeset | 245 | "enum = map (\<lambda>ys. the o map_of (zip (enum\<Colon>'a list) ys)) (List.n_lists (length (enum\<Colon>'a\<Colon>enum list)) enum)" | 
| 26444 | 246 | |
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 247 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 248 | "enum_all P = all_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 249 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 250 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 251 | "enum_ex P = ex_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 252 | |
| 26444 | 253 | instance proof | 
| 254 |   show "UNIV = set (enum \<Colon> ('a \<Rightarrow> 'b) list)"
 | |
| 255 | proof (rule UNIV_eq_I) | |
| 256 | fix f :: "'a \<Rightarrow> 'b" | |
| 257 | have "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))" | |
| 40683 | 258 | by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) | 
| 26444 | 259 | then show "f \<in> set enum" | 
| 40683 | 260 | by (auto simp add: enum_fun_def set_n_lists intro: in_enum) | 
| 26444 | 261 | qed | 
| 262 | next | |
| 263 | from map_of_zip_enum_inject | |
| 264 |   show "distinct (enum \<Colon> ('a \<Rightarrow> 'b) list)"
 | |
| 265 | by (auto intro!: inj_onI simp add: enum_fun_def | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 266 | distinct_map distinct_n_lists enum_distinct set_n_lists) | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 267 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 268 | fix P | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 269 |   show "enum_all (P :: ('a \<Rightarrow> 'b) \<Rightarrow> bool) = Ball UNIV P"
 | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 270 | proof | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 271 | assume "enum_all P" | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 272 | show "Ball UNIV P" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 273 | proof | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 274 | fix f :: "'a \<Rightarrow> 'b" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 275 | have f: "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 276 | by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 277 | from `enum_all P` have "P (the \<circ> map_of (zip enum (map f enum)))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 278 | unfolding enum_all_fun_def all_n_lists_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 279 | apply (simp add: set_n_lists) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 280 | apply (erule_tac x="map f enum" in allE) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 281 | apply (auto intro!: in_enum) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 282 | done | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 283 | from this f show "P f" by auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 284 | qed | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 285 | next | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 286 | assume "Ball UNIV P" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 287 | from this show "enum_all P" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 288 | unfolding enum_all_fun_def all_n_lists_def by auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 289 | qed | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 290 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 291 | fix P | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 292 |   show "enum_ex (P :: ('a \<Rightarrow> 'b) \<Rightarrow> bool) = Bex UNIV P"
 | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 293 | proof | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 294 | assume "enum_ex P" | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 295 | from this show "Bex UNIV P" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 296 | unfolding enum_ex_fun_def ex_n_lists_def by auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 297 | next | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 298 | assume "Bex UNIV P" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 299 | from this obtain f where "P f" .. | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 300 | have f: "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 301 | by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 302 | from `P f` this have "P (the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum)))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 303 | by auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 304 | from this show "enum_ex P" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 305 | unfolding enum_ex_fun_def ex_n_lists_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 306 | apply (auto simp add: set_n_lists) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 307 | apply (rule_tac x="map f enum" in exI) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 308 | apply (auto intro!: in_enum) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 309 | done | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 310 | qed | 
| 26444 | 311 | qed | 
| 312 | ||
| 313 | end | |
| 314 | ||
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 315 | lemma enum_fun_code [code]: "enum = (let enum_a = (enum \<Colon> 'a\<Colon>{enum, equal} list)
 | 
| 49948 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
48123diff
changeset | 316 | in map (\<lambda>ys. the o map_of (zip enum_a ys)) (List.n_lists (length enum_a) enum))" | 
| 28245 
9767dd8e1e54
celver code lemma avoid type ambiguity problem with Haskell
 haftmann parents: 
27487diff
changeset | 317 | by (simp add: enum_fun_def Let_def) | 
| 26444 | 318 | |
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 319 | lemma enum_all_fun_code [code]: | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 320 |   "enum_all P = (let enum_a = (enum :: 'a::{enum, equal} list)
 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 321 | in all_n_lists (\<lambda>bs. P (the o map_of (zip enum_a bs))) (length enum_a))" | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 322 | by (simp only: enum_all_fun_def Let_def) | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 323 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 324 | lemma enum_ex_fun_code [code]: | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 325 |   "enum_ex P = (let enum_a = (enum :: 'a::{enum, equal} list)
 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 326 | in ex_n_lists (\<lambda>bs. P (the o map_of (zip enum_a bs))) (length enum_a))" | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 327 | by (simp only: enum_ex_fun_def Let_def) | 
| 45963 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 328 | |
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 329 | instantiation set :: (enum) enum | 
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 330 | begin | 
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 331 | |
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 332 | definition | 
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 333 | "enum = map set (sublists enum)" | 
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 334 | |
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 335 | definition | 
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 336 | "enum_all P \<longleftrightarrow> (\<forall>A\<in>set enum. P (A::'a set))" | 
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 337 | |
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 338 | definition | 
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 339 | "enum_ex P \<longleftrightarrow> (\<exists>A\<in>set enum. P (A::'a set))" | 
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 340 | |
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 341 | instance proof | 
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 342 | qed (simp_all add: enum_set_def enum_all_set_def enum_ex_set_def sublists_powset distinct_set_sublists | 
| 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 343 | enum_distinct enum_UNIV) | 
| 29024 | 344 | |
| 345 | end | |
| 346 | ||
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 347 | instantiation unit :: enum | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 348 | begin | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 349 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 350 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 351 | "enum = [()]" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 352 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 353 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 354 | "enum_all P = P ()" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 355 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 356 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 357 | "enum_ex P = P ()" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 358 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 359 | instance proof | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 360 | qed (auto simp add: enum_unit_def enum_all_unit_def enum_ex_unit_def) | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 361 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 362 | end | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 363 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 364 | instantiation bool :: enum | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 365 | begin | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 366 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 367 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 368 | "enum = [False, True]" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 369 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 370 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 371 | "enum_all P \<longleftrightarrow> P False \<and> P True" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 372 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 373 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 374 | "enum_ex P \<longleftrightarrow> P False \<or> P True" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 375 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 376 | instance proof | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 377 | qed (simp_all only: enum_bool_def enum_all_bool_def enum_ex_bool_def UNIV_bool, simp_all) | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 378 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 379 | end | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 380 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 381 | instantiation prod :: (enum, enum) enum | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 382 | begin | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 383 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 384 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 385 | "enum = List.product enum enum" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 386 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 387 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 388 | "enum_all P = enum_all (%x. enum_all (%y. P (x, y)))" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 389 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 390 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 391 | "enum_ex P = enum_ex (%x. enum_ex (%y. P (x, y)))" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 392 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 393 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 394 | instance by default | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 395 | (simp_all add: enum_prod_def product_list_set distinct_product | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 396 | enum_UNIV enum_distinct enum_all_prod_def enum_ex_prod_def) | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 397 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 398 | end | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 399 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 400 | instantiation sum :: (enum, enum) enum | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 401 | begin | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 402 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 403 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 404 | "enum = map Inl enum @ map Inr enum" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 405 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 406 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 407 | "enum_all P \<longleftrightarrow> enum_all (\<lambda>x. P (Inl x)) \<and> enum_all (\<lambda>x. P (Inr x))" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 408 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 409 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 410 | "enum_ex P \<longleftrightarrow> enum_ex (\<lambda>x. P (Inl x)) \<or> enum_ex (\<lambda>x. P (Inr x))" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 411 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 412 | instance proof | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 413 | qed (simp_all only: enum_sum_def enum_all_sum_def enum_ex_sum_def UNIV_sum, | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 414 | auto simp add: enum_UNIV distinct_map enum_distinct) | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 415 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 416 | end | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 417 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 418 | instantiation option :: (enum) enum | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 419 | begin | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 420 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 421 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 422 | "enum = None # map Some enum" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 423 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 424 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 425 | "enum_all P \<longleftrightarrow> P None \<and> enum_all (\<lambda>x. P (Some x))" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 426 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 427 | definition | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 428 | "enum_ex P \<longleftrightarrow> P None \<or> enum_ex (\<lambda>x. P (Some x))" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 429 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 430 | instance proof | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 431 | qed (simp_all only: enum_option_def enum_all_option_def enum_ex_option_def UNIV_option_conv, | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 432 | auto simp add: distinct_map enum_UNIV enum_distinct) | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 433 | |
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 434 | end | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 435 | |
| 45963 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 haftmann parents: 
45144diff
changeset | 436 | |
| 40647 | 437 | subsection {* Small finite types *}
 | 
| 438 | ||
| 439 | text {* We define small finite types for the use in Quickcheck *}
 | |
| 440 | ||
| 441 | datatype finite_1 = a\<^isub>1 | |
| 442 | ||
| 40900 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 443 | notation (output) a\<^isub>1  ("a\<^isub>1")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 444 | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 445 | lemma UNIV_finite_1: | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 446 |   "UNIV = {a\<^isub>1}"
 | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 447 | by (auto intro: finite_1.exhaust) | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 448 | |
| 40647 | 449 | instantiation finite_1 :: enum | 
| 450 | begin | |
| 451 | ||
| 452 | definition | |
| 453 | "enum = [a\<^isub>1]" | |
| 454 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 455 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 456 | "enum_all P = P a\<^isub>1" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 457 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 458 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 459 | "enum_ex P = P a\<^isub>1" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 460 | |
| 40647 | 461 | instance proof | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 462 | qed (simp_all only: enum_finite_1_def enum_all_finite_1_def enum_ex_finite_1_def UNIV_finite_1, simp_all) | 
| 40647 | 463 | |
| 29024 | 464 | end | 
| 40647 | 465 | |
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 466 | instantiation finite_1 :: linorder | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 467 | begin | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 468 | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 469 | definition less_finite_1 :: "finite_1 \<Rightarrow> finite_1 \<Rightarrow> bool" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 470 | where | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 471 | "x < (y :: finite_1) \<longleftrightarrow> False" | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 472 | |
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 473 | definition less_eq_finite_1 :: "finite_1 \<Rightarrow> finite_1 \<Rightarrow> bool" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 474 | where | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 475 | "x \<le> (y :: finite_1) \<longleftrightarrow> True" | 
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 476 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 477 | instance | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 478 | apply (intro_classes) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 479 | apply (auto simp add: less_finite_1_def less_eq_finite_1_def) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 480 | apply (metis finite_1.exhaust) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 481 | done | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 482 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 483 | end | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 484 | |
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
41078diff
changeset | 485 | hide_const (open) a\<^isub>1 | 
| 40657 | 486 | |
| 40647 | 487 | datatype finite_2 = a\<^isub>1 | a\<^isub>2 | 
| 488 | ||
| 40900 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 489 | notation (output) a\<^isub>1  ("a\<^isub>1")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 490 | notation (output) a\<^isub>2  ("a\<^isub>2")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 491 | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 492 | lemma UNIV_finite_2: | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 493 |   "UNIV = {a\<^isub>1, a\<^isub>2}"
 | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 494 | by (auto intro: finite_2.exhaust) | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 495 | |
| 40647 | 496 | instantiation finite_2 :: enum | 
| 497 | begin | |
| 498 | ||
| 499 | definition | |
| 500 | "enum = [a\<^isub>1, a\<^isub>2]" | |
| 501 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 502 | definition | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 503 | "enum_all P \<longleftrightarrow> P a\<^isub>1 \<and> P a\<^isub>2" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 504 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 505 | definition | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 506 | "enum_ex P \<longleftrightarrow> P a\<^isub>1 \<or> P a\<^isub>2" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 507 | |
| 40647 | 508 | instance proof | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 509 | qed (simp_all only: enum_finite_2_def enum_all_finite_2_def enum_ex_finite_2_def UNIV_finite_2, simp_all) | 
| 40647 | 510 | |
| 511 | end | |
| 512 | ||
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 513 | instantiation finite_2 :: linorder | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 514 | begin | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 515 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 516 | definition less_finite_2 :: "finite_2 \<Rightarrow> finite_2 \<Rightarrow> bool" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 517 | where | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 518 | "x < y \<longleftrightarrow> x = a\<^isub>1 \<and> y = a\<^isub>2" | 
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 519 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 520 | definition less_eq_finite_2 :: "finite_2 \<Rightarrow> finite_2 \<Rightarrow> bool" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 521 | where | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 522 | "x \<le> y \<longleftrightarrow> x = y \<or> x < (y :: finite_2)" | 
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 523 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 524 | instance | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 525 | apply (intro_classes) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 526 | apply (auto simp add: less_finite_2_def less_eq_finite_2_def) | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 527 | apply (metis finite_2.nchotomy)+ | 
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 528 | done | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 529 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 530 | end | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 531 | |
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
41078diff
changeset | 532 | hide_const (open) a\<^isub>1 a\<^isub>2 | 
| 40657 | 533 | |
| 40647 | 534 | datatype finite_3 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 | 
| 535 | ||
| 40900 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 536 | notation (output) a\<^isub>1  ("a\<^isub>1")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 537 | notation (output) a\<^isub>2  ("a\<^isub>2")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 538 | notation (output) a\<^isub>3  ("a\<^isub>3")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 539 | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 540 | lemma UNIV_finite_3: | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 541 |   "UNIV = {a\<^isub>1, a\<^isub>2, a\<^isub>3}"
 | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 542 | by (auto intro: finite_3.exhaust) | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 543 | |
| 40647 | 544 | instantiation finite_3 :: enum | 
| 545 | begin | |
| 546 | ||
| 547 | definition | |
| 548 | "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3]" | |
| 549 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 550 | definition | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 551 | "enum_all P \<longleftrightarrow> P a\<^isub>1 \<and> P a\<^isub>2 \<and> P a\<^isub>3" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 552 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 553 | definition | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 554 | "enum_ex P \<longleftrightarrow> P a\<^isub>1 \<or> P a\<^isub>2 \<or> P a\<^isub>3" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 555 | |
| 40647 | 556 | instance proof | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 557 | qed (simp_all only: enum_finite_3_def enum_all_finite_3_def enum_ex_finite_3_def UNIV_finite_3, simp_all) | 
| 40647 | 558 | |
| 559 | end | |
| 560 | ||
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 561 | instantiation finite_3 :: linorder | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 562 | begin | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 563 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 564 | definition less_finite_3 :: "finite_3 \<Rightarrow> finite_3 \<Rightarrow> bool" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 565 | where | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 566 | "x < y = (case x of a\<^isub>1 \<Rightarrow> y \<noteq> a\<^isub>1 | a\<^isub>2 \<Rightarrow> y = a\<^isub>3 | a\<^isub>3 \<Rightarrow> False)" | 
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 567 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 568 | definition less_eq_finite_3 :: "finite_3 \<Rightarrow> finite_3 \<Rightarrow> bool" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 569 | where | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 570 | "x \<le> y \<longleftrightarrow> x = y \<or> x < (y :: finite_3)" | 
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 571 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 572 | instance proof (intro_classes) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 573 | qed (auto simp add: less_finite_3_def less_eq_finite_3_def split: finite_3.split_asm) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 574 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 575 | end | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 576 | |
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
41078diff
changeset | 577 | hide_const (open) a\<^isub>1 a\<^isub>2 a\<^isub>3 | 
| 40657 | 578 | |
| 40647 | 579 | datatype finite_4 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 | a\<^isub>4 | 
| 580 | ||
| 40900 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 581 | notation (output) a\<^isub>1  ("a\<^isub>1")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 582 | notation (output) a\<^isub>2  ("a\<^isub>2")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 583 | notation (output) a\<^isub>3  ("a\<^isub>3")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 584 | notation (output) a\<^isub>4  ("a\<^isub>4")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 585 | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 586 | lemma UNIV_finite_4: | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 587 |   "UNIV = {a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4}"
 | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 588 | by (auto intro: finite_4.exhaust) | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 589 | |
| 40647 | 590 | instantiation finite_4 :: enum | 
| 591 | begin | |
| 592 | ||
| 593 | definition | |
| 594 | "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4]" | |
| 595 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 596 | definition | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 597 | "enum_all P \<longleftrightarrow> P a\<^isub>1 \<and> P a\<^isub>2 \<and> P a\<^isub>3 \<and> P a\<^isub>4" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 598 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 599 | definition | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 600 | "enum_ex P \<longleftrightarrow> P a\<^isub>1 \<or> P a\<^isub>2 \<or> P a\<^isub>3 \<or> P a\<^isub>4" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 601 | |
| 40647 | 602 | instance proof | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 603 | qed (simp_all only: enum_finite_4_def enum_all_finite_4_def enum_ex_finite_4_def UNIV_finite_4, simp_all) | 
| 40647 | 604 | |
| 605 | end | |
| 606 | ||
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
41078diff
changeset | 607 | hide_const (open) a\<^isub>1 a\<^isub>2 a\<^isub>3 a\<^isub>4 | 
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 608 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 609 | |
| 40647 | 610 | datatype finite_5 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 | a\<^isub>4 | a\<^isub>5 | 
| 611 | ||
| 40900 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 612 | notation (output) a\<^isub>1  ("a\<^isub>1")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 613 | notation (output) a\<^isub>2  ("a\<^isub>2")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 614 | notation (output) a\<^isub>3  ("a\<^isub>3")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 615 | notation (output) a\<^isub>4  ("a\<^isub>4")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 616 | notation (output) a\<^isub>5  ("a\<^isub>5")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 617 | |
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 618 | lemma UNIV_finite_5: | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 619 |   "UNIV = {a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4, a\<^isub>5}"
 | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 620 | by (auto intro: finite_5.exhaust) | 
| 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 621 | |
| 40647 | 622 | instantiation finite_5 :: enum | 
| 623 | begin | |
| 624 | ||
| 625 | definition | |
| 626 | "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4, a\<^isub>5]" | |
| 627 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 628 | definition | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 629 | "enum_all P \<longleftrightarrow> P a\<^isub>1 \<and> P a\<^isub>2 \<and> P a\<^isub>3 \<and> P a\<^isub>4 \<and> P a\<^isub>5" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 630 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 631 | definition | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 632 | "enum_ex P \<longleftrightarrow> P a\<^isub>1 \<or> P a\<^isub>2 \<or> P a\<^isub>3 \<or> P a\<^isub>4 \<or> P a\<^isub>5" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 633 | |
| 40647 | 634 | instance proof | 
| 49950 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 haftmann parents: 
49949diff
changeset | 635 | qed (simp_all only: enum_finite_5_def enum_all_finite_5_def enum_ex_finite_5_def UNIV_finite_5, simp_all) | 
| 40647 | 636 | |
| 637 | end | |
| 638 | ||
| 46352 
73b03235799a
an executable version of accessible part (only for finite types yet)
 bulwahn parents: 
46336diff
changeset | 639 | hide_const (open) a\<^isub>1 a\<^isub>2 a\<^isub>3 a\<^isub>4 a\<^isub>5 | 
| 
73b03235799a
an executable version of accessible part (only for finite types yet)
 bulwahn parents: 
46336diff
changeset | 640 | |
| 49948 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
48123diff
changeset | 641 | |
| 46352 
73b03235799a
an executable version of accessible part (only for finite types yet)
 bulwahn parents: 
46336diff
changeset | 642 | subsection {* Closing up *}
 | 
| 40657 | 643 | |
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
41078diff
changeset | 644 | hide_type (open) finite_1 finite_2 finite_3 finite_4 finite_5 | 
| 49948 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
48123diff
changeset | 645 | hide_const (open) enum enum_all enum_ex all_n_lists ex_n_lists ntrancl | 
| 40647 | 646 | |
| 647 | end | |
| 49948 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
48123diff
changeset | 648 |