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