| author | haftmann | 
| Sat, 13 Apr 2019 08:11:48 +0000 | |
| changeset 70146 | 9839da71621f | 
| parent 69850 | 5f993636ac07 | 
| child 70523 | 1182ea5c9a6e | 
| permissions | -rw-r--r-- | 
| 31596 | 1  | 
(* Author: Florian Haftmann, TU Muenchen *)  | 
| 26348 | 2  | 
|
| 60758 | 3  | 
section \<open>Finite types as explicit enumerations\<close>  | 
| 26348 | 4  | 
|
5  | 
theory Enum  | 
|
| 58101 | 6  | 
imports Map Groups_List  | 
| 26348 | 7  | 
begin  | 
8  | 
||
| 61799 | 9  | 
subsection \<open>Class \<open>enum\<close>\<close>  | 
| 26348 | 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"  | 
| 61799 | 19  | 
\<comment> \<open>tailored towards simple instantiation\<close>  | 
| 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  | 
||
| 69593 | 55  | 
subsection \<open>Implementations using \<^class>\<open>enum\<close>\<close>  | 
| 49949 | 56  | 
|
| 60758 | 57  | 
subsubsection \<open>Unbounded operations and quantifiers\<close>  | 
| 49949 | 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]:  | 
| 67613 | 64  | 
"f -` B = set (filter (\<lambda>x. f x \<in> B) enum_class.enum)"  | 
| 
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
 | 
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  | 
||
| 60758 | 85  | 
subsubsection \<open>An executable choice operator\<close>  | 
| 49949 | 86  | 
|
87  | 
definition  | 
|
88  | 
[code del]: "enum_the = The"  | 
|
89  | 
||
90  | 
lemma [code]:  | 
|
| 67091 | 91  | 
"The P = (case filter P enum of [x] \<Rightarrow> x | _ \<Rightarrow> enum_the P)"  | 
| 49949 | 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])  | 
|
| 60758 | 109  | 
with \<open>P x\<close> in_enum[of x, unfolded enum_eq] \<open>x \<noteq> a\<close> show "False" by auto  | 
| 49949 | 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  | 
||
| 
54890
 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 
haftmann 
parents: 
54295 
diff
changeset
 | 
119  | 
declare [[code abort: enum_the]]  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
50567 
diff
changeset
 | 
120  | 
|
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
50567 
diff
changeset
 | 
121  | 
code_printing  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
50567 
diff
changeset
 | 
122  | 
constant enum_the \<rightharpoonup> (Eval) "(fn '_ => raise Match)"  | 
| 49949 | 123  | 
|
124  | 
||
| 60758 | 125  | 
subsubsection \<open>Equality and order on functions\<close>  | 
| 26348 | 126  | 
|
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37765 
diff
changeset
 | 
127  | 
instantiation "fun" :: (enum, equal) equal  | 
| 26513 | 128  | 
begin  | 
| 26348 | 129  | 
|
| 26513 | 130  | 
definition  | 
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37765 
diff
changeset
 | 
131  | 
"HOL.equal f g \<longleftrightarrow> (\<forall>x \<in> set enum. f x = g x)"  | 
| 26513 | 132  | 
|
| 31464 | 133  | 
instance proof  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
134  | 
qed (simp_all add: equal_fun_def fun_eq_iff enum_UNIV)  | 
| 26513 | 135  | 
|
136  | 
end  | 
|
| 26348 | 137  | 
|
| 
40898
 
882e860a1e83
changed order of lemmas to overwrite the general code equation with the nbe-specific one
 
bulwahn 
parents: 
40683 
diff
changeset
 | 
138  | 
lemma [code]:  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
139  | 
"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
 | 
140  | 
by (auto simp add: equal fun_eq_iff)  | 
| 
40898
 
882e860a1e83
changed order of lemmas to overwrite the general code equation with the nbe-specific one
 
bulwahn 
parents: 
40683 
diff
changeset
 | 
141  | 
|
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37765 
diff
changeset
 | 
142  | 
lemma [code nbe]:  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37765 
diff
changeset
 | 
143  | 
"HOL.equal (f :: _ \<Rightarrow> _) f \<longleftrightarrow> True"  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37765 
diff
changeset
 | 
144  | 
by (fact equal_refl)  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37765 
diff
changeset
 | 
145  | 
|
| 28562 | 146  | 
lemma order_fun [code]:  | 
| 61076 | 147  | 
fixes f g :: "'a::enum \<Rightarrow> 'b::order"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
148  | 
shows "f \<le> g \<longleftrightarrow> enum_all (\<lambda>x. f x \<le> g x)"  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
149  | 
and "f < g \<longleftrightarrow> f \<le> g \<and> enum_ex (\<lambda>x. f x \<noteq> g x)"  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
150  | 
by (simp_all add: fun_eq_iff le_fun_def order_less_le)  | 
| 26968 | 151  | 
|
152  | 
||
| 60758 | 153  | 
subsubsection \<open>Operations on relations\<close>  | 
| 49949 | 154  | 
|
155  | 
lemma [code]:  | 
|
156  | 
"Id = image (\<lambda>x. (x, x)) (set Enum.enum)"  | 
|
157  | 
by (auto intro: imageI in_enum)  | 
|
| 26968 | 158  | 
|
| 54148 | 159  | 
lemma tranclp_unfold [code]:  | 
| 49949 | 160  | 
  "tranclp r a b \<longleftrightarrow> (a, b) \<in> trancl {(x, y). r x y}"
 | 
161  | 
by (simp add: trancl_def)  | 
|
162  | 
||
| 54148 | 163  | 
lemma rtranclp_rtrancl_eq [code]:  | 
| 49949 | 164  | 
  "rtranclp r x y \<longleftrightarrow> (x, y) \<in> rtrancl {(x, y). r x y}"
 | 
165  | 
by (simp add: rtrancl_def)  | 
|
| 26968 | 166  | 
|
| 49949 | 167  | 
lemma max_ext_eq [code]:  | 
168  | 
  "max_ext R = {(X, Y). finite X \<and> finite Y \<and> Y \<noteq> {} \<and> (\<forall>x. x \<in> X \<longrightarrow> (\<exists>xa \<in> Y. (x, xa) \<in> R))}"
 | 
|
169  | 
by (auto simp add: max_ext.simps)  | 
|
170  | 
||
171  | 
lemma max_extp_eq [code]:  | 
|
172  | 
  "max_extp r x y \<longleftrightarrow> (x, y) \<in> max_ext {(x, y). r x y}"
 | 
|
173  | 
by (simp add: max_ext_def)  | 
|
| 26348 | 174  | 
|
| 49949 | 175  | 
lemma mlex_eq [code]:  | 
176  | 
  "f <*mlex*> R = {(x, y). f x < f y \<or> (f x \<le> f y \<and> (x, y) \<in> R)}"
 | 
|
177  | 
by (auto simp add: mlex_prod_def)  | 
|
178  | 
||
| 
55088
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
179  | 
|
| 60758 | 180  | 
subsubsection \<open>Bounded accessible part\<close>  | 
| 
55088
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
181  | 
|
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
182  | 
primrec bacc :: "('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> 'a set" 
 | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
183  | 
where  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
184  | 
  "bacc r 0 = {x. \<forall> y. (y, x) \<notin> r}"
 | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
185  | 
| "bacc r (Suc n) = (bacc r n \<union> {x. \<forall>y. (y, x) \<in> r \<longrightarrow> y \<in> bacc r n})"
 | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
186  | 
|
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
187  | 
lemma bacc_subseteq_acc:  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
188  | 
"bacc r n \<subseteq> Wellfounded.acc r"  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
189  | 
by (induct n) (auto intro: acc.intros)  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
190  | 
|
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
191  | 
lemma bacc_mono:  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
192  | 
"n \<le> m \<Longrightarrow> bacc r n \<subseteq> bacc r m"  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
193  | 
by (induct rule: dec_induct) auto  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
194  | 
|
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
195  | 
lemma bacc_upper_bound:  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
196  | 
  "bacc (r :: ('a \<times> 'a) set)  (card (UNIV :: 'a::finite set)) = (\<Union>n. bacc r n)"
 | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
197  | 
proof -  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
198  | 
have "mono (bacc r)" unfolding mono_def by (simp add: bacc_mono)  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
199  | 
moreover have "\<forall>n. bacc r n = bacc r (Suc n) \<longrightarrow> bacc r (Suc n) = bacc r (Suc (Suc n))" by auto  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
200  | 
moreover have "finite (range (bacc r))" by auto  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
201  | 
ultimately show ?thesis  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
202  | 
by (intro finite_mono_strict_prefix_implies_finite_fixpoint)  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
203  | 
(auto intro: finite_mono_remains_stable_implies_strict_prefix)  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
204  | 
qed  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
205  | 
|
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
206  | 
lemma acc_subseteq_bacc:  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
207  | 
assumes "finite r"  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
208  | 
shows "Wellfounded.acc r \<subseteq> (\<Union>n. bacc r n)"  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
209  | 
proof  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
210  | 
fix x  | 
| 67613 | 211  | 
assume "x \<in> Wellfounded.acc r"  | 
212  | 
then have "\<exists>n. x \<in> bacc r n"  | 
|
| 
55088
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
213  | 
proof (induct x arbitrary: rule: acc.induct)  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
214  | 
case (accI x)  | 
| 67091 | 215  | 
then have "\<forall>y. \<exists> n. (y, x) \<in> r \<longrightarrow> y \<in> bacc r n" by simp  | 
| 
55088
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
216  | 
from choice[OF this] obtain n where n: "\<forall>y. (y, x) \<in> r \<longrightarrow> y \<in> bacc r (n y)" ..  | 
| 67091 | 217  | 
obtain n where "\<And>y. (y, x) \<in> r \<Longrightarrow> y \<in> bacc r n"  | 
| 
55088
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
218  | 
proof  | 
| 67091 | 219  | 
fix y assume y: "(y, x) \<in> r"  | 
| 67613 | 220  | 
with n have "y \<in> bacc r (n y)" by auto  | 
| 67091 | 221  | 
moreover have "n y <= Max ((\<lambda>(y, x). n y) ` r)"  | 
| 60758 | 222  | 
using y \<open>finite r\<close> by (auto intro!: Max_ge)  | 
| 
55088
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
223  | 
note bacc_mono[OF this, of r]  | 
| 67613 | 224  | 
ultimately show "y \<in> bacc r (Max ((\<lambda>(y, x). n y) ` r))" by auto  | 
| 
55088
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
225  | 
qed  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
226  | 
then show ?case  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
227  | 
by (auto simp add: Let_def intro!: exI[of _ "Suc n"])  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
228  | 
qed  | 
| 67613 | 229  | 
then show "x \<in> (\<Union>n. bacc r n)" by auto  | 
| 
55088
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
230  | 
qed  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
231  | 
|
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
232  | 
lemma acc_bacc_eq:  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
233  | 
  fixes A :: "('a :: finite \<times> 'a) set"
 | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
234  | 
assumes "finite A"  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
235  | 
shows "Wellfounded.acc A = bacc A (card (UNIV :: 'a set))"  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
236  | 
using assms by (metis acc_subseteq_bacc bacc_subseteq_acc bacc_upper_bound order_eq_iff)  | 
| 
 
57c82e01022b
moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
 
blanchet 
parents: 
54890 
diff
changeset
 | 
237  | 
|
| 49949 | 238  | 
lemma [code]:  | 
239  | 
  fixes xs :: "('a::finite \<times> 'a) list"
 | 
|
| 54295 | 240  | 
  shows "Wellfounded.acc (set xs) = bacc (set xs) (card_UNIV TYPE('a))"
 | 
| 49949 | 241  | 
by (simp add: card_UNIV_def acc_bacc_eq)  | 
242  | 
||
| 26348 | 243  | 
|
| 69593 | 244  | 
subsection \<open>Default instances for \<^class>\<open>enum\<close>\<close>  | 
| 26348 | 245  | 
|
| 26444 | 246  | 
lemma map_of_zip_enum_is_Some:  | 
| 61076 | 247  | 
assumes "length ys = length (enum :: 'a::enum list)"  | 
248  | 
shows "\<exists>y. map_of (zip (enum :: 'a::enum list) ys) x = Some y"  | 
|
| 26444 | 249  | 
proof -  | 
| 61076 | 250  | 
from assms have "x \<in> set (enum :: 'a::enum list) \<longleftrightarrow>  | 
251  | 
(\<exists>y. map_of (zip (enum :: 'a::enum list) ys) x = Some y)"  | 
|
| 26444 | 252  | 
by (auto intro!: map_of_zip_is_Some)  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
253  | 
then show ?thesis using enum_UNIV by auto  | 
| 26444 | 254  | 
qed  | 
255  | 
||
256  | 
lemma map_of_zip_enum_inject:  | 
|
| 61076 | 257  | 
fixes xs ys :: "'b::enum list"  | 
258  | 
assumes length: "length xs = length (enum :: 'a::enum list)"  | 
|
259  | 
"length ys = length (enum :: 'a::enum list)"  | 
|
260  | 
and map_of: "the \<circ> map_of (zip (enum :: 'a::enum list) xs) = the \<circ> map_of (zip (enum :: 'a::enum list) ys)"  | 
|
| 26444 | 261  | 
shows "xs = ys"  | 
262  | 
proof -  | 
|
| 61076 | 263  | 
have "map_of (zip (enum :: 'a list) xs) = map_of (zip (enum :: 'a list) ys)"  | 
| 26444 | 264  | 
proof  | 
265  | 
fix x :: 'a  | 
|
266  | 
from length map_of_zip_enum_is_Some obtain y1 y2  | 
|
| 61076 | 267  | 
where "map_of (zip (enum :: 'a list) xs) x = Some y1"  | 
268  | 
and "map_of (zip (enum :: 'a list) ys) x = Some y2" by blast  | 
|
| 47230 | 269  | 
moreover from map_of  | 
| 61076 | 270  | 
have "the (map_of (zip (enum :: 'a::enum list) xs) x) = the (map_of (zip (enum :: 'a::enum list) ys) x)"  | 
| 26444 | 271  | 
by (auto dest: fun_cong)  | 
| 61076 | 272  | 
ultimately show "map_of (zip (enum :: 'a::enum list) xs) x = map_of (zip (enum :: 'a::enum list) ys) x"  | 
| 26444 | 273  | 
by simp  | 
274  | 
qed  | 
|
275  | 
with length enum_distinct show "xs = ys" by (rule map_of_zip_inject)  | 
|
276  | 
qed  | 
|
277  | 
||
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
278  | 
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
 | 
279  | 
where  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
280  | 
"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
 | 
281  | 
|
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
282  | 
lemma [code]:  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
283  | 
"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
 | 
284  | 
unfolding all_n_lists_def enum_all  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
285  | 
by (cases n) (auto simp add: enum_UNIV)  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
286  | 
|
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
287  | 
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
 | 
288  | 
where  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
289  | 
"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
 | 
290  | 
|
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
291  | 
lemma [code]:  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
292  | 
"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
 | 
293  | 
unfolding ex_n_lists_def enum_ex  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
294  | 
by (cases n) (auto simp add: enum_UNIV)  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
295  | 
|
| 26444 | 296  | 
instantiation "fun" :: (enum, enum) enum  | 
297  | 
begin  | 
|
298  | 
||
299  | 
definition  | 
|
| 67091 | 300  | 
"enum = map (\<lambda>ys. the \<circ> map_of (zip (enum::'a list) ys)) (List.n_lists (length (enum::'a::enum list)) enum)"  | 
| 26444 | 301  | 
|
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
302  | 
definition  | 
| 67091 | 303  | 
"enum_all P = all_n_lists (\<lambda>bs. P (the \<circ> map_of (zip enum bs))) (length (enum :: 'a list))"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
304  | 
|
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
305  | 
definition  | 
| 67091 | 306  | 
"enum_ex P = ex_n_lists (\<lambda>bs. P (the \<circ> map_of (zip enum bs))) (length (enum :: 'a list))"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
307  | 
|
| 26444 | 308  | 
instance proof  | 
| 61076 | 309  | 
  show "UNIV = set (enum :: ('a \<Rightarrow> 'b) list)"
 | 
| 26444 | 310  | 
proof (rule UNIV_eq_I)  | 
311  | 
fix f :: "'a \<Rightarrow> 'b"  | 
|
| 61076 | 312  | 
have "f = the \<circ> map_of (zip (enum :: 'a::enum list) (map f enum))"  | 
| 40683 | 313  | 
by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum)  | 
| 26444 | 314  | 
then show "f \<in> set enum"  | 
| 40683 | 315  | 
by (auto simp add: enum_fun_def set_n_lists intro: in_enum)  | 
| 26444 | 316  | 
qed  | 
317  | 
next  | 
|
318  | 
from map_of_zip_enum_inject  | 
|
| 61076 | 319  | 
  show "distinct (enum :: ('a \<Rightarrow> 'b) list)"
 | 
| 26444 | 320  | 
by (auto intro!: inj_onI simp add: enum_fun_def  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
321  | 
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
 | 
322  | 
next  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
323  | 
fix P  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
324  | 
  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
 | 
325  | 
proof  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
326  | 
assume "enum_all P"  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
327  | 
show "Ball UNIV P"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
328  | 
proof  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
329  | 
fix f :: "'a \<Rightarrow> 'b"  | 
| 61076 | 330  | 
have f: "f = the \<circ> map_of (zip (enum :: 'a::enum list) (map f enum))"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
331  | 
by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum)  | 
| 60758 | 332  | 
from \<open>enum_all P\<close> have "P (the \<circ> map_of (zip enum (map f enum)))"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
333  | 
unfolding enum_all_fun_def all_n_lists_def  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
334  | 
apply (simp add: set_n_lists)  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
335  | 
apply (erule_tac x="map f enum" in allE)  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
336  | 
apply (auto intro!: in_enum)  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
337  | 
done  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
338  | 
from this f show "P f" by auto  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
339  | 
qed  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
340  | 
next  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
341  | 
assume "Ball UNIV P"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
342  | 
from this show "enum_all P"  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
343  | 
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
 | 
344  | 
qed  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
345  | 
next  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
346  | 
fix P  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
347  | 
  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
 | 
348  | 
proof  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
349  | 
assume "enum_ex P"  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
350  | 
from this show "Bex UNIV P"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
351  | 
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
 | 
352  | 
next  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
353  | 
assume "Bex UNIV P"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
354  | 
from this obtain f where "P f" ..  | 
| 61076 | 355  | 
have f: "f = the \<circ> map_of (zip (enum :: 'a::enum list) (map f enum))"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
356  | 
by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum)  | 
| 61076 | 357  | 
from \<open>P f\<close> this have "P (the \<circ> map_of (zip (enum :: 'a::enum list) (map f enum)))"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
358  | 
by auto  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
359  | 
from this show "enum_ex P"  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
360  | 
unfolding enum_ex_fun_def ex_n_lists_def  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
361  | 
apply (auto simp add: set_n_lists)  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
362  | 
apply (rule_tac x="map f enum" in exI)  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
363  | 
apply (auto intro!: in_enum)  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
364  | 
done  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
365  | 
qed  | 
| 26444 | 366  | 
qed  | 
367  | 
||
368  | 
end  | 
|
369  | 
||
| 61076 | 370  | 
lemma enum_fun_code [code]: "enum = (let enum_a = (enum :: 'a::{enum, equal} list)
 | 
| 67091 | 371  | 
in map (\<lambda>ys. the \<circ> 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
 | 
372  | 
by (simp add: enum_fun_def Let_def)  | 
| 26444 | 373  | 
|
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
374  | 
lemma enum_all_fun_code [code]:  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
375  | 
  "enum_all P = (let enum_a = (enum :: 'a::{enum, equal} list)
 | 
| 67091 | 376  | 
in all_n_lists (\<lambda>bs. P (the \<circ> map_of (zip enum_a bs))) (length enum_a))"  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
377  | 
by (simp only: enum_all_fun_def Let_def)  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
378  | 
|
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
379  | 
lemma enum_ex_fun_code [code]:  | 
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
380  | 
  "enum_ex P = (let enum_a = (enum :: 'a::{enum, equal} list)
 | 
| 67091 | 381  | 
in ex_n_lists (\<lambda>bs. P (the \<circ> map_of (zip enum_a bs))) (length enum_a))"  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
382  | 
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
 | 
383  | 
|
| 
 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 
haftmann 
parents: 
45144 
diff
changeset
 | 
384  | 
instantiation set :: (enum) enum  | 
| 
 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 
haftmann 
parents: 
45144 
diff
changeset
 | 
385  | 
begin  | 
| 
 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 
haftmann 
parents: 
45144 
diff
changeset
 | 
386  | 
|
| 
 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 
haftmann 
parents: 
45144 
diff
changeset
 | 
387  | 
definition  | 
| 
65956
 
639eb3617a86
reorganised material on sublists
 
eberlm <eberlm@in.tum.de> 
parents: 
64592 
diff
changeset
 | 
388  | 
"enum = map set (subseqs enum)"  | 
| 
45963
 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 
haftmann 
parents: 
45144 
diff
changeset
 | 
389  | 
|
| 
 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 
haftmann 
parents: 
45144 
diff
changeset
 | 
390  | 
definition  | 
| 
 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 
haftmann 
parents: 
45144 
diff
changeset
 | 
391  | 
"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
 | 
392  | 
|
| 
 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 
haftmann 
parents: 
45144 
diff
changeset
 | 
393  | 
definition  | 
| 
 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 
haftmann 
parents: 
45144 
diff
changeset
 | 
394  | 
"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
 | 
395  | 
|
| 
 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 
haftmann 
parents: 
45144 
diff
changeset
 | 
396  | 
instance proof  | 
| 
65956
 
639eb3617a86
reorganised material on sublists
 
eberlm <eberlm@in.tum.de> 
parents: 
64592 
diff
changeset
 | 
397  | 
qed (simp_all add: enum_set_def enum_all_set_def enum_ex_set_def subseqs_powset distinct_set_subseqs  | 
| 
45963
 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 
haftmann 
parents: 
45144 
diff
changeset
 | 
398  | 
enum_distinct enum_UNIV)  | 
| 29024 | 399  | 
|
400  | 
end  | 
|
401  | 
||
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
402  | 
instantiation unit :: enum  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
403  | 
begin  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
404  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
405  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
406  | 
"enum = [()]"  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
407  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
408  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
409  | 
"enum_all P = P ()"  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
410  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
411  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
412  | 
"enum_ex P = P ()"  | 
| 
 
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  | 
instance proof  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
415  | 
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
 | 
416  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
417  | 
end  | 
| 
 
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  | 
instantiation bool :: enum  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
420  | 
begin  | 
| 
 
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 = [False, True]"  | 
| 
 
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_all P \<longleftrightarrow> P False \<and> P True"  | 
| 
 
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  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
429  | 
"enum_ex P \<longleftrightarrow> P False \<or> P True"  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
430  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
431  | 
instance proof  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
432  | 
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
 | 
433  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
434  | 
end  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
435  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
436  | 
instantiation prod :: (enum, enum) enum  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
437  | 
begin  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
438  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
439  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
440  | 
"enum = List.product enum enum"  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
441  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
442  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
443  | 
"enum_all P = enum_all (%x. enum_all (%y. P (x, y)))"  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
444  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
445  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
446  | 
"enum_ex P = enum_ex (%x. enum_ex (%y. P (x, y)))"  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
447  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
448  | 
|
| 61169 | 449  | 
instance  | 
450  | 
by standard  | 
|
451  | 
(simp_all add: enum_prod_def distinct_product  | 
|
452  | 
enum_UNIV enum_distinct enum_all_prod_def enum_ex_prod_def)  | 
|
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
453  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
454  | 
end  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
455  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
456  | 
instantiation sum :: (enum, enum) enum  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
457  | 
begin  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
458  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
459  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
460  | 
"enum = map Inl enum @ map Inr enum"  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
461  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
462  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
463  | 
"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
 | 
464  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
465  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
466  | 
"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
 | 
467  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
468  | 
instance proof  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
469  | 
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
 | 
470  | 
auto simp add: enum_UNIV distinct_map enum_distinct)  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
471  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
472  | 
end  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
473  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
474  | 
instantiation option :: (enum) enum  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
475  | 
begin  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
476  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
477  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
478  | 
"enum = None # map Some enum"  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
479  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
480  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
481  | 
"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
 | 
482  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
483  | 
definition  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
484  | 
"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
 | 
485  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
486  | 
instance proof  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
487  | 
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
 | 
488  | 
auto simp add: distinct_map enum_UNIV enum_distinct)  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
489  | 
|
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
490  | 
end  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
491  | 
|
| 
45963
 
1c7e6454883e
enum type class instance for `set`; dropped misfitting code lemma for trancl
 
haftmann 
parents: 
45144 
diff
changeset
 | 
492  | 
|
| 60758 | 493  | 
subsection \<open>Small finite types\<close>  | 
| 40647 | 494  | 
|
| 60758 | 495  | 
text \<open>We define small finite types for use in Quickcheck\<close>  | 
| 40647 | 496  | 
|
| 
58659
 
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
 
wenzelm 
parents: 
58646 
diff
changeset
 | 
497  | 
datatype (plugins only: code "quickcheck" extraction) finite_1 =  | 
| 
58350
 
919149921e46
added 'extraction' plugins -- this might help 'HOL-Proofs'
 
blanchet 
parents: 
58334 
diff
changeset
 | 
498  | 
a\<^sub>1  | 
| 40647 | 499  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
500  | 
notation (output) a\<^sub>1  ("a\<^sub>1")
 | 
| 
40900
 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 
bulwahn 
parents: 
40898 
diff
changeset
 | 
501  | 
|
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
502  | 
lemma UNIV_finite_1:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
503  | 
  "UNIV = {a\<^sub>1}"
 | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
504  | 
by (auto intro: finite_1.exhaust)  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
505  | 
|
| 40647 | 506  | 
instantiation finite_1 :: enum  | 
507  | 
begin  | 
|
508  | 
||
509  | 
definition  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
510  | 
"enum = [a\<^sub>1]"  | 
| 40647 | 511  | 
|
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
512  | 
definition  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
513  | 
"enum_all P = P a\<^sub>1"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
514  | 
|
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
515  | 
definition  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
516  | 
"enum_ex P = P a\<^sub>1"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
517  | 
|
| 40647 | 518  | 
instance proof  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
519  | 
qed (simp_all only: enum_finite_1_def enum_all_finite_1_def enum_ex_finite_1_def UNIV_finite_1, simp_all)  | 
| 40647 | 520  | 
|
| 29024 | 521  | 
end  | 
| 40647 | 522  | 
|
| 
40651
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
523  | 
instantiation finite_1 :: linorder  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
524  | 
begin  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
525  | 
|
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
526  | 
definition less_finite_1 :: "finite_1 \<Rightarrow> finite_1 \<Rightarrow> bool"  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
527  | 
where  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
528  | 
"x < (y :: finite_1) \<longleftrightarrow> False"  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
529  | 
|
| 
40651
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
530  | 
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
 | 
531  | 
where  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
532  | 
"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
 | 
533  | 
|
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
534  | 
instance  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
535  | 
apply (intro_classes)  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
536  | 
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
 | 
537  | 
apply (metis finite_1.exhaust)  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
538  | 
done  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
539  | 
|
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
540  | 
end  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
541  | 
|
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
542  | 
instance finite_1 :: "{dense_linorder, wellorder}"
 | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
543  | 
by intro_classes (simp_all add: less_finite_1_def)  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
544  | 
|
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
545  | 
instantiation finite_1 :: complete_lattice  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
546  | 
begin  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
547  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
548  | 
definition [simp]: "Inf = (\<lambda>_. a\<^sub>1)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
549  | 
definition [simp]: "Sup = (\<lambda>_. a\<^sub>1)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
550  | 
definition [simp]: "bot = a\<^sub>1"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
551  | 
definition [simp]: "top = a\<^sub>1"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
552  | 
definition [simp]: "inf = (\<lambda>_ _. a\<^sub>1)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
553  | 
definition [simp]: "sup = (\<lambda>_ _. a\<^sub>1)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
554  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
555  | 
instance by intro_classes(simp_all add: less_eq_finite_1_def)  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
556  | 
end  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
557  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
558  | 
instance finite_1 :: complete_distrib_lattice  | 
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
61799 
diff
changeset
 | 
559  | 
by standard simp_all  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
560  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
561  | 
instance finite_1 :: complete_linorder ..  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
562  | 
|
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
563  | 
lemma finite_1_eq: "x = a\<^sub>1"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
564  | 
by(cases x) simp  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
565  | 
|
| 60758 | 566  | 
simproc_setup finite_1_eq ("x::finite_1") = \<open>
 | 
| 59582 | 567  | 
fn _ => fn _ => fn ct =>  | 
568  | 
(case Thm.term_of ct of  | 
|
| 69593 | 569  | 
Const (\<^const_name>\<open>a\<^sub>1\<close>, _) => NONE  | 
| 59582 | 570  | 
    | _ => SOME (mk_meta_eq @{thm finite_1_eq}))
 | 
| 60758 | 571  | 
\<close>  | 
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
572  | 
|
| 59582 | 573  | 
instantiation finite_1 :: complete_boolean_algebra  | 
574  | 
begin  | 
|
| 67399 | 575  | 
definition [simp]: "(-) = (\<lambda>_ _. a\<^sub>1)"  | 
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
576  | 
definition [simp]: "uminus = (\<lambda>_. a\<^sub>1)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
577  | 
instance by intro_classes simp_all  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
578  | 
end  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
579  | 
|
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
580  | 
instantiation finite_1 ::  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
581  | 
  "{linordered_ring_strict, linordered_comm_semiring_strict, ordered_comm_ring,
 | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
582  | 
ordered_cancel_comm_monoid_diff, comm_monoid_mult, ordered_ring_abs,  | 
| 64290 | 583  | 
one, modulo, sgn, inverse}"  | 
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
584  | 
begin  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
585  | 
definition [simp]: "Groups.zero = a\<^sub>1"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
586  | 
definition [simp]: "Groups.one = a\<^sub>1"  | 
| 67399 | 587  | 
definition [simp]: "(+) = (\<lambda>_ _. a\<^sub>1)"  | 
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
67951 
diff
changeset
 | 
588  | 
definition [simp]: "(*) = (\<lambda>_ _. a\<^sub>1)"  | 
| 67399 | 589  | 
definition [simp]: "(mod) = (\<lambda>_ _. a\<^sub>1)"  | 
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
590  | 
definition [simp]: "abs = (\<lambda>_. a\<^sub>1)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
591  | 
definition [simp]: "sgn = (\<lambda>_. a\<^sub>1)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
592  | 
definition [simp]: "inverse = (\<lambda>_. a\<^sub>1)"  | 
| 
60352
 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 
haftmann 
parents: 
59867 
diff
changeset
 | 
593  | 
definition [simp]: "divide = (\<lambda>_ _. a\<^sub>1)"  | 
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
594  | 
|
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
595  | 
instance by intro_classes(simp_all add: less_finite_1_def)  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
596  | 
end  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
597  | 
|
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
598  | 
declare [[simproc del: finite_1_eq]]  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
599  | 
hide_const (open) a\<^sub>1  | 
| 40657 | 600  | 
|
| 
58659
 
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
 
wenzelm 
parents: 
58646 
diff
changeset
 | 
601  | 
datatype (plugins only: code "quickcheck" extraction) finite_2 =  | 
| 
58350
 
919149921e46
added 'extraction' plugins -- this might help 'HOL-Proofs'
 
blanchet 
parents: 
58334 
diff
changeset
 | 
602  | 
a\<^sub>1 | a\<^sub>2  | 
| 40647 | 603  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
604  | 
notation (output) a\<^sub>1  ("a\<^sub>1")
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
605  | 
notation (output) a\<^sub>2  ("a\<^sub>2")
 | 
| 
40900
 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 
bulwahn 
parents: 
40898 
diff
changeset
 | 
606  | 
|
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
607  | 
lemma UNIV_finite_2:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
608  | 
  "UNIV = {a\<^sub>1, a\<^sub>2}"
 | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
609  | 
by (auto intro: finite_2.exhaust)  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
610  | 
|
| 40647 | 611  | 
instantiation finite_2 :: enum  | 
612  | 
begin  | 
|
613  | 
||
614  | 
definition  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
615  | 
"enum = [a\<^sub>1, a\<^sub>2]"  | 
| 40647 | 616  | 
|
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
617  | 
definition  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
618  | 
"enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
619  | 
|
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
620  | 
definition  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
621  | 
"enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
622  | 
|
| 40647 | 623  | 
instance proof  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
624  | 
qed (simp_all only: enum_finite_2_def enum_all_finite_2_def enum_ex_finite_2_def UNIV_finite_2, simp_all)  | 
| 40647 | 625  | 
|
626  | 
end  | 
|
627  | 
||
| 
40651
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
628  | 
instantiation finite_2 :: linorder  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
629  | 
begin  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
630  | 
|
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
631  | 
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
 | 
632  | 
where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
633  | 
"x < y \<longleftrightarrow> x = a\<^sub>1 \<and> y = a\<^sub>2"  | 
| 
40651
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
634  | 
|
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
635  | 
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
 | 
636  | 
where  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
637  | 
"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
 | 
638  | 
|
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
639  | 
instance  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
640  | 
apply (intro_classes)  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
641  | 
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
 | 
642  | 
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
 | 
643  | 
done  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
644  | 
|
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
645  | 
end  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
646  | 
|
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
647  | 
instance finite_2 :: wellorder  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
648  | 
by(rule wf_wellorderI)(simp add: less_finite_2_def, intro_classes)  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
649  | 
|
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
650  | 
instantiation finite_2 :: complete_lattice  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
651  | 
begin  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
652  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
653  | 
definition "\<Sqinter>A = (if a\<^sub>1 \<in> A then a\<^sub>1 else a\<^sub>2)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
654  | 
definition "\<Squnion>A = (if a\<^sub>2 \<in> A then a\<^sub>2 else a\<^sub>1)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
655  | 
definition [simp]: "bot = a\<^sub>1"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
656  | 
definition [simp]: "top = a\<^sub>2"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
657  | 
definition "x \<sqinter> y = (if x = a\<^sub>1 \<or> y = a\<^sub>1 then a\<^sub>1 else a\<^sub>2)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
658  | 
definition "x \<squnion> y = (if x = a\<^sub>2 \<or> y = a\<^sub>2 then a\<^sub>2 else a\<^sub>1)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
659  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
660  | 
lemma neq_finite_2_a\<^sub>1_iff [simp]: "x \<noteq> a\<^sub>1 \<longleftrightarrow> x = a\<^sub>2"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
661  | 
by(cases x) simp_all  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
662  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
663  | 
lemma neq_finite_2_a\<^sub>1_iff' [simp]: "a\<^sub>1 \<noteq> x \<longleftrightarrow> x = a\<^sub>2"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
664  | 
by(cases x) simp_all  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
665  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
666  | 
lemma neq_finite_2_a\<^sub>2_iff [simp]: "x \<noteq> a\<^sub>2 \<longleftrightarrow> x = a\<^sub>1"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
667  | 
by(cases x) simp_all  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
668  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
669  | 
lemma neq_finite_2_a\<^sub>2_iff' [simp]: "a\<^sub>2 \<noteq> x \<longleftrightarrow> x = a\<^sub>1"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
670  | 
by(cases x) simp_all  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
671  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
672  | 
instance  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
673  | 
proof  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
674  | 
fix x :: finite_2 and A  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
675  | 
assume "x \<in> A"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
676  | 
then show "\<Sqinter>A \<le> x" "x \<le> \<Squnion>A"  | 
| 69850 | 677  | 
by(cases x; auto simp add: less_eq_finite_2_def less_finite_2_def Inf_finite_2_def Sup_finite_2_def)+  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
678  | 
qed(auto simp add: less_eq_finite_2_def less_finite_2_def inf_finite_2_def sup_finite_2_def Inf_finite_2_def Sup_finite_2_def)  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
679  | 
end  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
680  | 
|
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
681  | 
instance finite_2 :: complete_linorder ..  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
682  | 
|
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
683  | 
instance finite_2 :: complete_distrib_lattice ..  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
684  | 
|
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
685  | 
instantiation finite_2 :: "{field, idom_abs_sgn, idom_modulo}" begin
 | 
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
686  | 
definition [simp]: "0 = a\<^sub>1"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
687  | 
definition [simp]: "1 = a\<^sub>2"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
688  | 
definition "x + y = (case (x, y) of (a\<^sub>1, a\<^sub>1) \<Rightarrow> a\<^sub>1 | (a\<^sub>2, a\<^sub>2) \<Rightarrow> a\<^sub>1 | _ \<Rightarrow> a\<^sub>2)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
689  | 
definition "uminus = (\<lambda>x :: finite_2. x)"  | 
| 67399 | 690  | 
definition "(-) = ((+) :: finite_2 \<Rightarrow> _)"  | 
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
691  | 
definition "x * y = (case (x, y) of (a\<^sub>2, a\<^sub>2) \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> a\<^sub>1)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
692  | 
definition "inverse = (\<lambda>x :: finite_2. x)"  | 
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
67951 
diff
changeset
 | 
693  | 
definition "divide = ((*) :: finite_2 \<Rightarrow> _)"  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
694  | 
definition "x mod y = (case (x, y) of (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> a\<^sub>1)"  | 
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
695  | 
definition "abs = (\<lambda>x :: finite_2. x)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
696  | 
definition "sgn = (\<lambda>x :: finite_2. x)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
697  | 
instance  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
698  | 
by standard  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
699  | 
(simp_all add: plus_finite_2_def uminus_finite_2_def minus_finite_2_def  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
700  | 
times_finite_2_def  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
701  | 
inverse_finite_2_def divide_finite_2_def modulo_finite_2_def  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
702  | 
abs_finite_2_def sgn_finite_2_def  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
703  | 
split: finite_2.splits)  | 
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
704  | 
end  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
705  | 
|
| 
58646
 
cd63a4b12a33
specialized specification: avoid trivial instances
 
haftmann 
parents: 
58350 
diff
changeset
 | 
706  | 
lemma two_finite_2 [simp]:  | 
| 
 
cd63a4b12a33
specialized specification: avoid trivial instances
 
haftmann 
parents: 
58350 
diff
changeset
 | 
707  | 
"2 = a\<^sub>1"  | 
| 
 
cd63a4b12a33
specialized specification: avoid trivial instances
 
haftmann 
parents: 
58350 
diff
changeset
 | 
708  | 
by (simp add: numeral.simps plus_finite_2_def)  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
709  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
710  | 
lemma dvd_finite_2_unfold:  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
711  | 
"x dvd y \<longleftrightarrow> x = a\<^sub>2 \<or> y = a\<^sub>1"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
712  | 
by (auto simp add: dvd_def times_finite_2_def split: finite_2.splits)  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
713  | 
|
| 66817 | 714  | 
instantiation finite_2 :: "{normalization_semidom, unique_euclidean_semiring}" begin
 | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
715  | 
definition [simp]: "normalize = (id :: finite_2 \<Rightarrow> _)"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
716  | 
definition [simp]: "unit_factor = (id :: finite_2 \<Rightarrow> _)"  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
717  | 
definition [simp]: "euclidean_size x = (case x of a\<^sub>1 \<Rightarrow> 0 | a\<^sub>2 \<Rightarrow> 1)"  | 
| 
66838
 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 
haftmann 
parents: 
66817 
diff
changeset
 | 
718  | 
definition [simp]: "division_segment (x :: finite_2) = 1"  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
719  | 
instance  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
720  | 
by standard  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
721  | 
(auto simp add: divide_finite_2_def times_finite_2_def dvd_finite_2_unfold  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
722  | 
split: finite_2.splits)  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
723  | 
end  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
724  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
725  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
726  | 
hide_const (open) a\<^sub>1 a\<^sub>2  | 
| 40657 | 727  | 
|
| 
58659
 
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
 
wenzelm 
parents: 
58646 
diff
changeset
 | 
728  | 
datatype (plugins only: code "quickcheck" extraction) finite_3 =  | 
| 
58350
 
919149921e46
added 'extraction' plugins -- this might help 'HOL-Proofs'
 
blanchet 
parents: 
58334 
diff
changeset
 | 
729  | 
a\<^sub>1 | a\<^sub>2 | a\<^sub>3  | 
| 40647 | 730  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
731  | 
notation (output) a\<^sub>1  ("a\<^sub>1")
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
732  | 
notation (output) a\<^sub>2  ("a\<^sub>2")
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
733  | 
notation (output) a\<^sub>3  ("a\<^sub>3")
 | 
| 
40900
 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 
bulwahn 
parents: 
40898 
diff
changeset
 | 
734  | 
|
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
735  | 
lemma UNIV_finite_3:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
736  | 
  "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3}"
 | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
737  | 
by (auto intro: finite_3.exhaust)  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
738  | 
|
| 40647 | 739  | 
instantiation finite_3 :: enum  | 
740  | 
begin  | 
|
741  | 
||
742  | 
definition  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
743  | 
"enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3]"  | 
| 40647 | 744  | 
|
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
745  | 
definition  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
746  | 
"enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2 \<and> P a\<^sub>3"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
747  | 
|
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
748  | 
definition  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
749  | 
"enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2 \<or> P a\<^sub>3"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
750  | 
|
| 40647 | 751  | 
instance proof  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
752  | 
qed (simp_all only: enum_finite_3_def enum_all_finite_3_def enum_ex_finite_3_def UNIV_finite_3, simp_all)  | 
| 40647 | 753  | 
|
754  | 
end  | 
|
755  | 
||
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
756  | 
lemma finite_3_not_eq_unfold:  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
757  | 
  "x \<noteq> a\<^sub>1 \<longleftrightarrow> x \<in> {a\<^sub>2, a\<^sub>3}"
 | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
758  | 
  "x \<noteq> a\<^sub>2 \<longleftrightarrow> x \<in> {a\<^sub>1, a\<^sub>3}"
 | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
759  | 
  "x \<noteq> a\<^sub>3 \<longleftrightarrow> x \<in> {a\<^sub>1, a\<^sub>2}"
 | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
760  | 
by (cases x; simp)+  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
761  | 
|
| 
40651
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
762  | 
instantiation finite_3 :: linorder  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
763  | 
begin  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
764  | 
|
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
765  | 
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
 | 
766  | 
where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
767  | 
"x < y = (case x of a\<^sub>1 \<Rightarrow> y \<noteq> a\<^sub>1 | a\<^sub>2 \<Rightarrow> y = a\<^sub>3 | a\<^sub>3 \<Rightarrow> False)"  | 
| 
40651
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
768  | 
|
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
769  | 
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
 | 
770  | 
where  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
771  | 
"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
 | 
772  | 
|
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
773  | 
instance proof (intro_classes)  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
774  | 
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
 | 
775  | 
|
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
776  | 
end  | 
| 
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
777  | 
|
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
778  | 
instance finite_3 :: wellorder  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
779  | 
proof(rule wf_wellorderI)  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
780  | 
  have "inv_image less_than (case_finite_3 0 1 2) = {(x, y). x < y}"
 | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
781  | 
by(auto simp add: less_finite_3_def split: finite_3.splits)  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
782  | 
from this[symmetric] show "wf \<dots>" by simp  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
783  | 
qed intro_classes  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
784  | 
|
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
785  | 
class finite_lattice = finite + lattice + Inf + Sup + bot + top +  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
786  | 
  assumes Inf_finite_empty: "Inf {} = Sup UNIV"
 | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
787  | 
assumes Inf_finite_insert: "Inf (insert a A) = a \<sqinter> Inf A"  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
788  | 
  assumes Sup_finite_empty: "Sup {} = Inf UNIV"
 | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
789  | 
assumes Sup_finite_insert: "Sup (insert a A) = a \<squnion> Sup A"  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
790  | 
assumes bot_finite_def: "bot = Inf UNIV"  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
791  | 
assumes top_finite_def: "top = Sup UNIV"  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
792  | 
begin  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
793  | 
|
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
794  | 
subclass complete_lattice  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
795  | 
proof  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
796  | 
fix x A  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
797  | 
show "x \<in> A \<Longrightarrow> \<Sqinter>A \<le> x"  | 
| 
67951
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
798  | 
by (metis Set.set_insert abel_semigroup.commute local.Inf_finite_insert local.inf.abel_semigroup_axioms local.inf.left_idem local.inf.orderI)  | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
799  | 
show "x \<in> A \<Longrightarrow> x \<le> \<Squnion>A"  | 
| 
67951
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
800  | 
by (metis Set.set_insert insert_absorb2 local.Sup_finite_insert local.sup.absorb_iff2)  | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
801  | 
next  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
802  | 
fix A z  | 
| 
67951
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
803  | 
have "\<Squnion> UNIV = z \<squnion> \<Squnion>UNIV"  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
804  | 
by (subst Sup_finite_insert [symmetric], simp add: insert_UNIV)  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
805  | 
from this have [simp]: "z \<le> \<Squnion>UNIV"  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
806  | 
using local.le_iff_sup by auto  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
807  | 
have "(\<forall> x. x \<in> A \<longrightarrow> z \<le> x) \<longrightarrow> z \<le> \<Sqinter>A"  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
808  | 
apply (rule finite_induct [of A "\<lambda> A . (\<forall> x. x \<in> A \<longrightarrow> z \<le> x) \<longrightarrow> z \<le> \<Sqinter>A"])  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
809  | 
by (simp_all add: Inf_finite_empty Inf_finite_insert)  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
810  | 
from this show "(\<And>x. x \<in> A \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> \<Sqinter>A"  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
811  | 
by simp  | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
812  | 
|
| 
67951
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
813  | 
have "\<Sqinter> UNIV = z \<sqinter> \<Sqinter>UNIV"  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
814  | 
by (subst Inf_finite_insert [symmetric], simp add: insert_UNIV)  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
815  | 
from this have [simp]: "\<Sqinter>UNIV \<le> z"  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
816  | 
by (simp add: local.inf.absorb_iff2)  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
817  | 
have "(\<forall> x. x \<in> A \<longrightarrow> x \<le> z) \<longrightarrow> \<Squnion>A \<le> z"  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
818  | 
by (rule finite_induct [of A "\<lambda> A . (\<forall> x. x \<in> A \<longrightarrow> x \<le> z) \<longrightarrow> \<Squnion>A \<le> z" ], simp_all add: Sup_finite_empty Sup_finite_insert)  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
819  | 
from this show " (\<And>x. x \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> \<Squnion>A \<le> z"  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
820  | 
by blast  | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
821  | 
next  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
822  | 
  show "\<Sqinter>{} = \<top>"
 | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
823  | 
by (simp add: Inf_finite_empty top_finite_def)  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
824  | 
  show " \<Squnion>{} = \<bottom>"
 | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
825  | 
by (simp add: Sup_finite_empty bot_finite_def)  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
826  | 
qed  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
827  | 
end  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
828  | 
|
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
829  | 
class finite_distrib_lattice = finite_lattice + distrib_lattice  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
830  | 
begin  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
831  | 
lemma finite_inf_Sup: "a \<sqinter> (Sup A) = Sup {a \<sqinter> b | b . b \<in> A}"
 | 
| 
67951
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
832  | 
proof (rule finite_induct [of A "\<lambda> A . a \<sqinter> (Sup A) = Sup {a \<sqinter> b | b . b \<in> A}"], simp_all)
 | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
833  | 
fix x::"'a"  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
834  | 
fix F  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
835  | 
assume "x \<notin> F"  | 
| 
67951
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
836  | 
  assume [simp]: "a \<sqinter> \<Squnion>F = \<Squnion>{a \<sqinter> b |b. b \<in> F}"
 | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
837  | 
  have [simp]: " insert (a \<sqinter> x) {a \<sqinter> b |b. b \<in> F} = {a \<sqinter> b |b. b = x \<or> b \<in> F}"
 | 
| 
67951
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
838  | 
by blast  | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
839  | 
have "a \<sqinter> (x \<squnion> \<Squnion>F) = a \<sqinter> x \<squnion> a \<sqinter> \<Squnion>F"  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
840  | 
by (simp add: inf_sup_distrib1)  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
841  | 
  also have "... = a \<sqinter> x \<squnion> \<Squnion>{a \<sqinter> b |b. b \<in> F}"
 | 
| 
67951
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
842  | 
by simp  | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
843  | 
  also have "... = \<Squnion>{a \<sqinter> b |b. b = x \<or> b \<in> F}"
 | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
844  | 
by (unfold Sup_insert[THEN sym], simp)  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
845  | 
  finally show "a \<sqinter> (x \<squnion> \<Squnion>F) = \<Squnion>{a \<sqinter> b |b. b = x \<or> b \<in> F}"
 | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
846  | 
by simp  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
847  | 
qed  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
848  | 
|
| 69275 | 849  | 
lemma finite_Inf_Sup: "\<Sqinter>(Sup ` A) \<le> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
 | 
850  | 
proof (rule finite_induct [of A "\<lambda>A. \<Sqinter>(Sup ` A) \<le> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"], simp_all add: finite_UnionD)
 | 
|
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
851  | 
fix x::"'a set"  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
852  | 
fix F  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
853  | 
assume "x \<notin> F"  | 
| 
67951
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
854  | 
  have [simp]: "{\<Squnion>x \<sqinter> b |b . b \<in> Inf ` {f ` F |f. \<forall>Y\<in>F. f Y \<in> Y} } = {\<Squnion>x \<sqinter> (Inf (f ` F)) |f  . (\<forall>Y\<in>F. f Y \<in> Y)}"
 | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
855  | 
by auto  | 
| 
67951
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
856  | 
define fa where "fa = (\<lambda> (b::'a) f Y . (if Y = x then b else f Y))"  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
857  | 
  have "\<And>f b. \<forall>Y\<in>F. f Y \<in> Y \<Longrightarrow> b \<in> x \<Longrightarrow> insert b (f ` (F \<inter> {Y. Y \<noteq> x})) = insert (fa b f x) (fa b f ` F) \<and> fa b f x \<in> x \<and> (\<forall>Y\<in>F. fa b f Y \<in> Y)"
 | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
858  | 
by (auto simp add: fa_def)  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
859  | 
  from this have B: "\<And>f b. \<forall>Y\<in>F. f Y \<in> Y \<Longrightarrow> b \<in> x \<Longrightarrow> fa b f ` ({x} \<union> F) \<in> {insert (f x) (f ` F) |f. f x \<in> x \<and> (\<forall>Y\<in>F. f Y \<in> Y)}"
 | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
860  | 
by blast  | 
| 69275 | 861  | 
  have [simp]: "\<And>f b. \<forall>Y\<in>F. f Y \<in> Y \<Longrightarrow> b \<in> x \<Longrightarrow> b \<sqinter> (\<Sqinter>x\<in>F. f x)  \<le> \<Squnion>(Inf ` {insert (f x) (f ` F) |f. f x \<in> x \<and> (\<forall>Y\<in>F. f Y \<in> Y)})"
 | 
| 69768 | 862  | 
using B apply (rule SUP_upper2)  | 
863  | 
using \<open>x \<notin> F\<close> apply (simp_all add: fa_def Inf_union_distrib)  | 
|
864  | 
apply (simp add: image_mono Inf_superset_mono inf.coboundedI2)  | 
|
865  | 
done  | 
|
| 69275 | 866  | 
  assume "\<Sqinter>(Sup ` F) \<le> \<Squnion>(Inf ` {f ` F |f. \<forall>Y\<in>F. f Y \<in> Y})"
 | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
867  | 
|
| 69275 | 868  | 
  from this have "\<Squnion>x \<sqinter> \<Sqinter>(Sup ` F) \<le> \<Squnion>x \<sqinter> \<Squnion>(Inf ` {f ` F |f. \<forall>Y\<in>F. f Y \<in> Y})"
 | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
869  | 
apply simp  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
870  | 
using inf.coboundedI2 by blast  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
871  | 
  also have "... = Sup {\<Squnion>x \<sqinter> (Inf (f ` F)) |f  .  (\<forall>Y\<in>F. f Y \<in> Y)}"
 | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
872  | 
by (simp add: finite_inf_Sup)  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
873  | 
|
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
874  | 
  also have "... = Sup {Sup {Inf (f ` F) \<sqinter> b | b . b \<in> x} |f  .  (\<forall>Y\<in>F. f Y \<in> Y)}"
 | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
875  | 
apply (subst inf_commute)  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
876  | 
by (simp add: finite_inf_Sup)  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
877  | 
|
| 69275 | 878  | 
  also have "... \<le> \<Squnion>(Inf ` {insert (f x) (f ` F) |f. f x \<in> x \<and> (\<forall>Y\<in>F. f Y \<in> Y)})"
 | 
| 
67951
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
879  | 
apply (rule Sup_least, clarsimp)+  | 
| 
 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67829 
diff
changeset
 | 
880  | 
by (subst inf_commute, simp)  | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
881  | 
|
| 69275 | 882  | 
  finally show "\<Squnion>x \<sqinter> \<Sqinter>(Sup ` F) \<le> \<Squnion>(Inf ` {insert (f x) (f ` F) |f. f x \<in> x \<and> (\<forall>Y\<in>F. f Y \<in> Y)})"
 | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
883  | 
by simp  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
884  | 
qed  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
885  | 
|
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
886  | 
subclass complete_distrib_lattice  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
887  | 
by (standard, rule finite_Inf_Sup)  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
888  | 
end  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
889  | 
|
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
890  | 
instantiation finite_3 :: finite_lattice  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
891  | 
begin  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
892  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
893  | 
definition "\<Sqinter>A = (if a\<^sub>1 \<in> A then a\<^sub>1 else if a\<^sub>2 \<in> A then a\<^sub>2 else a\<^sub>3)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
894  | 
definition "\<Squnion>A = (if a\<^sub>3 \<in> A then a\<^sub>3 else if a\<^sub>2 \<in> A then a\<^sub>2 else a\<^sub>1)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
895  | 
definition [simp]: "bot = a\<^sub>1"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
896  | 
definition [simp]: "top = a\<^sub>3"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
897  | 
definition [simp]: "inf = (min :: finite_3 \<Rightarrow> _)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
898  | 
definition [simp]: "sup = (max :: finite_3 \<Rightarrow> _)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
899  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
900  | 
instance  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
901  | 
proof  | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
902  | 
qed (auto simp add: Inf_finite_3_def Sup_finite_3_def max_def min_def less_eq_finite_3_def less_finite_3_def split: finite_3.split)  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
903  | 
end  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
904  | 
|
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
905  | 
instance finite_3 :: complete_lattice ..  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
906  | 
|
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
907  | 
instance finite_3 :: finite_distrib_lattice  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
908  | 
proof  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
909  | 
qed (auto simp add: min_def max_def)  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
910  | 
|
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
911  | 
instance finite_3 :: complete_distrib_lattice ..  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
912  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
913  | 
instance finite_3 :: complete_linorder ..  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
914  | 
|
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
915  | 
instantiation finite_3 :: "{field, idom_abs_sgn, idom_modulo}" begin
 | 
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
916  | 
definition [simp]: "0 = a\<^sub>1"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
917  | 
definition [simp]: "1 = a\<^sub>2"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
918  | 
definition  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
919  | 
"x + y = (case (x, y) of  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
920  | 
(a\<^sub>1, a\<^sub>1) \<Rightarrow> a\<^sub>1 | (a\<^sub>2, a\<^sub>3) \<Rightarrow> a\<^sub>1 | (a\<^sub>3, a\<^sub>2) \<Rightarrow> a\<^sub>1  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
921  | 
| (a\<^sub>1, a\<^sub>2) \<Rightarrow> a\<^sub>2 | (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | (a\<^sub>3, a\<^sub>3) \<Rightarrow> a\<^sub>2  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
922  | 
| _ \<Rightarrow> a\<^sub>3)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
923  | 
definition "- x = (case x of a\<^sub>1 \<Rightarrow> a\<^sub>1 | a\<^sub>2 \<Rightarrow> a\<^sub>3 | a\<^sub>3 \<Rightarrow> a\<^sub>2)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
924  | 
definition "x - y = x + (- y :: finite_3)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
925  | 
definition "x * y = (case (x, y) of (a\<^sub>2, a\<^sub>2) \<Rightarrow> a\<^sub>2 | (a\<^sub>3, a\<^sub>3) \<Rightarrow> a\<^sub>2 | (a\<^sub>2, a\<^sub>3) \<Rightarrow> a\<^sub>3 | (a\<^sub>3, a\<^sub>2) \<Rightarrow> a\<^sub>3 | _ \<Rightarrow> a\<^sub>1)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
926  | 
definition "inverse = (\<lambda>x :: finite_3. x)"  | 
| 
60429
 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 
haftmann 
parents: 
60352 
diff
changeset
 | 
927  | 
definition "x div y = x * inverse (y :: finite_3)"  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
928  | 
definition "x mod y = (case y of a\<^sub>1 \<Rightarrow> x | _ \<Rightarrow> a\<^sub>1)"  | 
| 64290 | 929  | 
definition "abs = (\<lambda>x. case x of a\<^sub>3 \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> x)"  | 
930  | 
definition "sgn = (\<lambda>x :: finite_3. x)"  | 
|
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
931  | 
instance  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
932  | 
by standard  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
933  | 
(simp_all add: plus_finite_3_def uminus_finite_3_def minus_finite_3_def  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
934  | 
times_finite_3_def  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
935  | 
inverse_finite_3_def divide_finite_3_def modulo_finite_3_def  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
936  | 
abs_finite_3_def sgn_finite_3_def  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
937  | 
less_finite_3_def  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
938  | 
split: finite_3.splits)  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
939  | 
end  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
940  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
941  | 
lemma two_finite_3 [simp]:  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
942  | 
"2 = a\<^sub>3"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
943  | 
by (simp add: numeral.simps plus_finite_3_def)  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
944  | 
|
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
945  | 
lemma dvd_finite_3_unfold:  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
946  | 
"x dvd y \<longleftrightarrow> x = a\<^sub>2 \<or> x = a\<^sub>3 \<or> y = a\<^sub>1"  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
947  | 
by (cases x) (auto simp add: dvd_def times_finite_3_def split: finite_3.splits)  | 
| 
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
948  | 
|
| 66817 | 949  | 
instantiation finite_3 :: "{normalization_semidom, unique_euclidean_semiring}" begin
 | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
950  | 
definition [simp]: "normalize x = (case x of a\<^sub>3 \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> x)"  | 
| 
64592
 
7759f1766189
more fine-grained type class hierarchy for div and mod
 
haftmann 
parents: 
64290 
diff
changeset
 | 
951  | 
definition [simp]: "unit_factor = (id :: finite_3 \<Rightarrow> _)"  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
952  | 
definition [simp]: "euclidean_size x = (case x of a\<^sub>1 \<Rightarrow> 0 | _ \<Rightarrow> 1)"  | 
| 
66838
 
17989f6bc7b2
clarified uniqueness criterion for euclidean rings
 
haftmann 
parents: 
66817 
diff
changeset
 | 
953  | 
definition [simp]: "division_segment (x :: finite_3) = 1"  | 
| 
66806
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
954  | 
instance proof  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
955  | 
fix x :: finite_3  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
956  | 
assume "x \<noteq> 0"  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
957  | 
then show "is_unit (unit_factor x)"  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
958  | 
by (cases x) (simp_all add: dvd_finite_3_unfold)  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
959  | 
qed (auto simp add: divide_finite_3_def times_finite_3_def  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
960  | 
dvd_finite_3_unfold inverse_finite_3_def plus_finite_3_def  | 
| 
 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 
haftmann 
parents: 
65956 
diff
changeset
 | 
961  | 
split: finite_3.splits)  | 
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
962  | 
end  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
963  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
964  | 
hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3  | 
| 40657 | 965  | 
|
| 
58659
 
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
 
wenzelm 
parents: 
58646 
diff
changeset
 | 
966  | 
datatype (plugins only: code "quickcheck" extraction) finite_4 =  | 
| 
58350
 
919149921e46
added 'extraction' plugins -- this might help 'HOL-Proofs'
 
blanchet 
parents: 
58334 
diff
changeset
 | 
967  | 
a\<^sub>1 | a\<^sub>2 | a\<^sub>3 | a\<^sub>4  | 
| 40647 | 968  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
969  | 
notation (output) a\<^sub>1  ("a\<^sub>1")
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
970  | 
notation (output) a\<^sub>2  ("a\<^sub>2")
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
971  | 
notation (output) a\<^sub>3  ("a\<^sub>3")
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
972  | 
notation (output) a\<^sub>4  ("a\<^sub>4")
 | 
| 
40900
 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 
bulwahn 
parents: 
40898 
diff
changeset
 | 
973  | 
|
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
974  | 
lemma UNIV_finite_4:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
975  | 
  "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4}"
 | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
976  | 
by (auto intro: finite_4.exhaust)  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
977  | 
|
| 40647 | 978  | 
instantiation finite_4 :: enum  | 
979  | 
begin  | 
|
980  | 
||
981  | 
definition  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
982  | 
"enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4]"  | 
| 40647 | 983  | 
|
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
984  | 
definition  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
985  | 
"enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2 \<and> P a\<^sub>3 \<and> P a\<^sub>4"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
986  | 
|
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
987  | 
definition  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
988  | 
"enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2 \<or> P a\<^sub>3 \<or> P a\<^sub>4"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
989  | 
|
| 40647 | 990  | 
instance proof  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
991  | 
qed (simp_all only: enum_finite_4_def enum_all_finite_4_def enum_ex_finite_4_def UNIV_finite_4, simp_all)  | 
| 40647 | 992  | 
|
993  | 
end  | 
|
994  | 
||
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
995  | 
instantiation finite_4 :: finite_distrib_lattice begin  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
996  | 
|
| 69593 | 997  | 
text \<open>\<^term>\<open>a\<^sub>1\<close> $<$ \<^term>\<open>a\<^sub>2\<close>,\<^term>\<open>a\<^sub>3\<close> $<$ \<^term>\<open>a\<^sub>4\<close>,  | 
998  | 
but \<^term>\<open>a\<^sub>2\<close> and \<^term>\<open>a\<^sub>3\<close> are incomparable.\<close>  | 
|
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
999  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1000  | 
definition  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1001  | 
"x < y \<longleftrightarrow> (case (x, y) of  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1002  | 
(a\<^sub>1, a\<^sub>1) \<Rightarrow> False | (a\<^sub>1, _) \<Rightarrow> True  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1003  | 
| (a\<^sub>2, a\<^sub>4) \<Rightarrow> True  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1004  | 
| (a\<^sub>3, a\<^sub>4) \<Rightarrow> True | _ \<Rightarrow> False)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1005  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1006  | 
definition  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1007  | 
"x \<le> y \<longleftrightarrow> (case (x, y) of  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1008  | 
(a\<^sub>1, _) \<Rightarrow> True  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1009  | 
| (a\<^sub>2, a\<^sub>2) \<Rightarrow> True | (a\<^sub>2, a\<^sub>4) \<Rightarrow> True  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1010  | 
| (a\<^sub>3, a\<^sub>3) \<Rightarrow> True | (a\<^sub>3, a\<^sub>4) \<Rightarrow> True  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1011  | 
| (a\<^sub>4, a\<^sub>4) \<Rightarrow> True | _ \<Rightarrow> False)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1012  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1013  | 
definition  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1014  | 
"\<Sqinter>A = (if a\<^sub>1 \<in> A \<or> a\<^sub>2 \<in> A \<and> a\<^sub>3 \<in> A then a\<^sub>1 else if a\<^sub>2 \<in> A then a\<^sub>2 else if a\<^sub>3 \<in> A then a\<^sub>3 else a\<^sub>4)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1015  | 
definition  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1016  | 
"\<Squnion>A = (if a\<^sub>4 \<in> A \<or> a\<^sub>2 \<in> A \<and> a\<^sub>3 \<in> A then a\<^sub>4 else if a\<^sub>2 \<in> A then a\<^sub>2 else if a\<^sub>3 \<in> A then a\<^sub>3 else a\<^sub>1)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1017  | 
definition [simp]: "bot = a\<^sub>1"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1018  | 
definition [simp]: "top = a\<^sub>4"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1019  | 
definition  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1020  | 
"x \<sqinter> y = (case (x, y) of  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1021  | 
(a\<^sub>1, _) \<Rightarrow> a\<^sub>1 | (_, a\<^sub>1) \<Rightarrow> a\<^sub>1 | (a\<^sub>2, a\<^sub>3) \<Rightarrow> a\<^sub>1 | (a\<^sub>3, a\<^sub>2) \<Rightarrow> a\<^sub>1  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1022  | 
| (a\<^sub>2, _) \<Rightarrow> a\<^sub>2 | (_, a\<^sub>2) \<Rightarrow> a\<^sub>2  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1023  | 
| (a\<^sub>3, _) \<Rightarrow> a\<^sub>3 | (_, a\<^sub>3) \<Rightarrow> a\<^sub>3  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1024  | 
| _ \<Rightarrow> a\<^sub>4)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1025  | 
definition  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1026  | 
"x \<squnion> y = (case (x, y) of  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1027  | 
(a\<^sub>4, _) \<Rightarrow> a\<^sub>4 | (_, a\<^sub>4) \<Rightarrow> a\<^sub>4 | (a\<^sub>2, a\<^sub>3) \<Rightarrow> a\<^sub>4 | (a\<^sub>3, a\<^sub>2) \<Rightarrow> a\<^sub>4  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1028  | 
| (a\<^sub>2, _) \<Rightarrow> a\<^sub>2 | (_, a\<^sub>2) \<Rightarrow> a\<^sub>2  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1029  | 
| (a\<^sub>3, _) \<Rightarrow> a\<^sub>3 | (_, a\<^sub>3) \<Rightarrow> a\<^sub>3  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1030  | 
| _ \<Rightarrow> a\<^sub>1)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1031  | 
|
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1032  | 
instance proof  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1033  | 
qed(auto simp add: less_finite_4_def less_eq_finite_4_def Inf_finite_4_def Sup_finite_4_def  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1034  | 
inf_finite_4_def sup_finite_4_def split: finite_4.splits)  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1035  | 
end  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1036  | 
|
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1037  | 
instance finite_4 :: complete_lattice ..  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1038  | 
|
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1039  | 
instance finite_4 :: complete_distrib_lattice ..  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1040  | 
|
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
1041  | 
instantiation finite_4 :: complete_boolean_algebra begin  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
1042  | 
definition "- x = (case x of a\<^sub>1 \<Rightarrow> a\<^sub>4 | a\<^sub>2 \<Rightarrow> a\<^sub>3 | a\<^sub>3 \<Rightarrow> a\<^sub>2 | a\<^sub>4 \<Rightarrow> a\<^sub>1)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
1043  | 
definition "x - y = x \<sqinter> - (y :: finite_4)"  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
1044  | 
instance  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
1045  | 
by intro_classes  | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1046  | 
(simp_all add: inf_finite_4_def sup_finite_4_def uminus_finite_4_def minus_finite_4_def  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1047  | 
split: finite_4.splits)  | 
| 
57922
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
1048  | 
end  | 
| 
 
dc78785427d0
add algebraic type class instances for Enum.finite* types
 
Andreas Lochbihler 
parents: 
57818 
diff
changeset
 | 
1049  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
1050  | 
hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4  | 
| 
40651
 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 
bulwahn 
parents: 
40650 
diff
changeset
 | 
1051  | 
|
| 
58659
 
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
 
wenzelm 
parents: 
58646 
diff
changeset
 | 
1052  | 
datatype (plugins only: code "quickcheck" extraction) finite_5 =  | 
| 
58350
 
919149921e46
added 'extraction' plugins -- this might help 'HOL-Proofs'
 
blanchet 
parents: 
58334 
diff
changeset
 | 
1053  | 
a\<^sub>1 | a\<^sub>2 | a\<^sub>3 | a\<^sub>4 | a\<^sub>5  | 
| 40647 | 1054  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
1055  | 
notation (output) a\<^sub>1  ("a\<^sub>1")
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
1056  | 
notation (output) a\<^sub>2  ("a\<^sub>2")
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
1057  | 
notation (output) a\<^sub>3  ("a\<^sub>3")
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
1058  | 
notation (output) a\<^sub>4  ("a\<^sub>4")
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
1059  | 
notation (output) a\<^sub>5  ("a\<^sub>5")
 | 
| 
40900
 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 
bulwahn 
parents: 
40898 
diff
changeset
 | 
1060  | 
|
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
1061  | 
lemma UNIV_finite_5:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
1062  | 
  "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4, a\<^sub>5}"
 | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
1063  | 
by (auto intro: finite_5.exhaust)  | 
| 
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
1064  | 
|
| 40647 | 1065  | 
instantiation finite_5 :: enum  | 
1066  | 
begin  | 
|
1067  | 
||
1068  | 
definition  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
1069  | 
"enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4, a\<^sub>5]"  | 
| 40647 | 1070  | 
|
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
1071  | 
definition  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
1072  | 
"enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2 \<and> P a\<^sub>3 \<and> P a\<^sub>4 \<and> P a\<^sub>5"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
1073  | 
|
| 
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
1074  | 
definition  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
1075  | 
"enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2 \<or> P a\<^sub>3 \<or> P a\<^sub>4 \<or> P a\<^sub>5"  | 
| 
41078
 
051251fde456
adding more efficient implementations for quantifiers in Enum
 
bulwahn 
parents: 
40900 
diff
changeset
 | 
1076  | 
|
| 40647 | 1077  | 
instance proof  | 
| 
49950
 
cd882d53ba6b
tailored enum specification towards simple instantiation;
 
haftmann 
parents: 
49949 
diff
changeset
 | 
1078  | 
qed (simp_all only: enum_finite_5_def enum_all_finite_5_def enum_ex_finite_5_def UNIV_finite_5, simp_all)  | 
| 40647 | 1079  | 
|
1080  | 
end  | 
|
1081  | 
||
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1082  | 
instantiation finite_5 :: finite_lattice  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1083  | 
begin  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1084  | 
|
| 60758 | 1085  | 
text \<open>The non-distributive pentagon lattice $N_5$\<close>  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1086  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1087  | 
definition  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1088  | 
"x < y \<longleftrightarrow> (case (x, y) of  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1089  | 
(a\<^sub>1, a\<^sub>1) \<Rightarrow> False | (a\<^sub>1, _) \<Rightarrow> True  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1090  | 
| (a\<^sub>2, a\<^sub>3) \<Rightarrow> True | (a\<^sub>2, a\<^sub>5) \<Rightarrow> True  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1091  | 
| (a\<^sub>3, a\<^sub>5) \<Rightarrow> True  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1092  | 
| (a\<^sub>4, a\<^sub>5) \<Rightarrow> True | _ \<Rightarrow> False)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1093  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1094  | 
definition  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1095  | 
"x \<le> y \<longleftrightarrow> (case (x, y) of  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1096  | 
(a\<^sub>1, _) \<Rightarrow> True  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1097  | 
| (a\<^sub>2, a\<^sub>2) \<Rightarrow> True | (a\<^sub>2, a\<^sub>3) \<Rightarrow> True | (a\<^sub>2, a\<^sub>5) \<Rightarrow> True  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1098  | 
| (a\<^sub>3, a\<^sub>3) \<Rightarrow> True | (a\<^sub>3, a\<^sub>5) \<Rightarrow> True  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1099  | 
| (a\<^sub>4, a\<^sub>4) \<Rightarrow> True | (a\<^sub>4, a\<^sub>5) \<Rightarrow> True  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1100  | 
| (a\<^sub>5, a\<^sub>5) \<Rightarrow> True | _ \<Rightarrow> False)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1101  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1102  | 
definition  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1103  | 
"\<Sqinter>A =  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1104  | 
(if a\<^sub>1 \<in> A \<or> a\<^sub>4 \<in> A \<and> (a\<^sub>2 \<in> A \<or> a\<^sub>3 \<in> A) then a\<^sub>1  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1105  | 
else if a\<^sub>2 \<in> A then a\<^sub>2  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1106  | 
else if a\<^sub>3 \<in> A then a\<^sub>3  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1107  | 
else if a\<^sub>4 \<in> A then a\<^sub>4  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1108  | 
else a\<^sub>5)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1109  | 
definition  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1110  | 
"\<Squnion>A =  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1111  | 
(if a\<^sub>5 \<in> A \<or> a\<^sub>4 \<in> A \<and> (a\<^sub>2 \<in> A \<or> a\<^sub>3 \<in> A) then a\<^sub>5  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1112  | 
else if a\<^sub>3 \<in> A then a\<^sub>3  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1113  | 
else if a\<^sub>2 \<in> A then a\<^sub>2  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1114  | 
else if a\<^sub>4 \<in> A then a\<^sub>4  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1115  | 
else a\<^sub>1)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1116  | 
definition [simp]: "bot = a\<^sub>1"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1117  | 
definition [simp]: "top = a\<^sub>5"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1118  | 
definition  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1119  | 
"x \<sqinter> y = (case (x, y) of  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1120  | 
(a\<^sub>1, _) \<Rightarrow> a\<^sub>1 | (_, a\<^sub>1) \<Rightarrow> a\<^sub>1 | (a\<^sub>2, a\<^sub>4) \<Rightarrow> a\<^sub>1 | (a\<^sub>4, a\<^sub>2) \<Rightarrow> a\<^sub>1 | (a\<^sub>3, a\<^sub>4) \<Rightarrow> a\<^sub>1 | (a\<^sub>4, a\<^sub>3) \<Rightarrow> a\<^sub>1  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1121  | 
| (a\<^sub>2, _) \<Rightarrow> a\<^sub>2 | (_, a\<^sub>2) \<Rightarrow> a\<^sub>2  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1122  | 
| (a\<^sub>3, _) \<Rightarrow> a\<^sub>3 | (_, a\<^sub>3) \<Rightarrow> a\<^sub>3  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1123  | 
| (a\<^sub>4, _) \<Rightarrow> a\<^sub>4 | (_, a\<^sub>4) \<Rightarrow> a\<^sub>4  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1124  | 
| _ \<Rightarrow> a\<^sub>5)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1125  | 
definition  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1126  | 
"x \<squnion> y = (case (x, y) of  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1127  | 
(a\<^sub>5, _) \<Rightarrow> a\<^sub>5 | (_, a\<^sub>5) \<Rightarrow> a\<^sub>5 | (a\<^sub>2, a\<^sub>4) \<Rightarrow> a\<^sub>5 | (a\<^sub>4, a\<^sub>2) \<Rightarrow> a\<^sub>5 | (a\<^sub>3, a\<^sub>4) \<Rightarrow> a\<^sub>5 | (a\<^sub>4, a\<^sub>3) \<Rightarrow> a\<^sub>5  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1128  | 
| (a\<^sub>3, _) \<Rightarrow> a\<^sub>3 | (_, a\<^sub>3) \<Rightarrow> a\<^sub>3  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1129  | 
| (a\<^sub>2, _) \<Rightarrow> a\<^sub>2 | (_, a\<^sub>2) \<Rightarrow> a\<^sub>2  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1130  | 
| (a\<^sub>4, _) \<Rightarrow> a\<^sub>4 | (_, a\<^sub>4) \<Rightarrow> a\<^sub>4  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1131  | 
| _ \<Rightarrow> a\<^sub>1)"  | 
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1132  | 
|
| 
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1133  | 
instance  | 
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1134  | 
proof  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1135  | 
qed (auto simp add: less_eq_finite_5_def less_finite_5_def inf_finite_5_def sup_finite_5_def  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1136  | 
Inf_finite_5_def Sup_finite_5_def split: finite_5.splits if_split_asm)  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1137  | 
end  | 
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1138  | 
|
| 
67829
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1139  | 
|
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1140  | 
instance finite_5 :: complete_lattice ..  | 
| 
 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
67613 
diff
changeset
 | 
1141  | 
|
| 
57818
 
51aa30c9ee4e
add complete_lattice instances for Enum.finite_* types such that quickcheck deals with lattice class operations
 
Andreas Lochbihler 
parents: 
57247 
diff
changeset
 | 
1142  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
1143  | 
hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4 a\<^sub>5  | 
| 
46352
 
73b03235799a
an executable version of accessible part (only for finite types yet)
 
bulwahn 
parents: 
46336 
diff
changeset
 | 
1144  | 
|
| 
49948
 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 
haftmann 
parents: 
48123 
diff
changeset
 | 
1145  | 
|
| 60758 | 1146  | 
subsection \<open>Closing up\<close>  | 
| 40657 | 1147  | 
|
| 
41085
 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 
bulwahn 
parents: 
41078 
diff
changeset
 | 
1148  | 
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
 | 
1149  | 
hide_const (open) enum enum_all enum_ex all_n_lists ex_n_lists ntrancl  | 
| 40647 | 1150  | 
|
1151  | 
end  |