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