1(* Author: Florian Haftmann, TU Muenchen *)
2
3section \<open>Finite types as explicit enumerations\<close>
4
5theory Enum
6imports Map Groups_List
7begin
8
9subsection \<open>Class \<open>enum\<close>\<close>
10
11class enum =
12  fixes enum :: "'a list"
13  fixes enum_all :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
14  fixes enum_ex :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
15  assumes UNIV_enum: "UNIV = set enum"
16    and enum_distinct: "distinct enum"
17  assumes enum_all_UNIV: "enum_all P \<longleftrightarrow> Ball UNIV P"
18  assumes enum_ex_UNIV: "enum_ex P \<longleftrightarrow> Bex UNIV P" 
19   \<comment> \<open>tailored towards simple instantiation\<close>
20begin
21
22subclass finite proof
23qed (simp add: UNIV_enum)
24
25lemma enum_UNIV:
26  "set enum = UNIV"
27  by (simp only: UNIV_enum)
28
29lemma in_enum: "x \<in> set enum"
30  by (simp add: enum_UNIV)
31
32lemma enum_eq_I:
33  assumes "\<And>x. x \<in> set xs"
34  shows "set enum = set xs"
35proof -
36  from assms UNIV_eq_I have "UNIV = set xs" by auto
37  with enum_UNIV show ?thesis by simp
38qed
39
40lemma card_UNIV_length_enum:
41  "card (UNIV :: 'a set) = length enum"
42  by (simp add: UNIV_enum distinct_card enum_distinct)
43
44lemma enum_all [simp]:
45  "enum_all = HOL.All"
46  by (simp add: fun_eq_iff enum_all_UNIV)
47
48lemma enum_ex [simp]:
49  "enum_ex = HOL.Ex" 
50  by (simp add: fun_eq_iff enum_ex_UNIV)
51
52end
53
54
55subsection \<open>Implementations using @{class enum}\<close>
56
57subsubsection \<open>Unbounded operations and quantifiers\<close>
58
59lemma Collect_code [code]:
60  "Collect P = set (filter P enum)"
61  by (simp add: enum_UNIV)
62
63lemma vimage_code [code]:
64  "f -` B = set (filter (\<lambda>x. f x \<in> B) enum_class.enum)"
65  unfolding vimage_def Collect_code ..
66
67definition card_UNIV :: "'a itself \<Rightarrow> nat"
68where
69  [code del]: "card_UNIV TYPE('a) = card (UNIV :: 'a set)"
70
71lemma [code]:
72  "card_UNIV TYPE('a :: enum) = card (set (Enum.enum :: 'a list))"
73  by (simp only: card_UNIV_def enum_UNIV)
74
75lemma all_code [code]: "(\<forall>x. P x) \<longleftrightarrow> enum_all P"
76  by simp
77
78lemma exists_code [code]: "(\<exists>x. P x) \<longleftrightarrow> enum_ex P"
79  by simp
80
81lemma exists1_code [code]: "(\<exists>!x. P x) \<longleftrightarrow> list_ex1 P enum"
82  by (auto simp add: list_ex1_iff enum_UNIV)
83
84
85subsubsection \<open>An executable choice operator\<close>
86
87definition
88  [code del]: "enum_the = The"
89
90lemma [code]:
91  "The P = (case filter P enum of [x] \<Rightarrow> x | _ \<Rightarrow> enum_the P)"
92proof -
93  {
94    fix a
95    assume filter_enum: "filter P enum = [a]"
96    have "The P = a"
97    proof (rule the_equality)
98      fix x
99      assume "P x"
100      show "x = a"
101      proof (rule ccontr)
102        assume "x \<noteq> a"
103        from filter_enum obtain us vs
104          where enum_eq: "enum = us @ [a] @ vs"
105          and "\<forall> x \<in> set us. \<not> P x"
106          and "\<forall> x \<in> set vs. \<not> P x"
107          and "P a"
108          by (auto simp add: filter_eq_Cons_iff) (simp only: filter_empty_conv[symmetric])
109        with \<open>P x\<close> in_enum[of x, unfolded enum_eq] \<open>x \<noteq> a\<close> show "False" by auto
110      qed
111    next
112      from filter_enum show "P a" by (auto simp add: filter_eq_Cons_iff)
113    qed
114  }
115  from this show ?thesis
116    unfolding enum_the_def by (auto split: list.split)
117qed
118
119declare [[code abort: enum_the]]
120
121code_printing
122  constant enum_the \<rightharpoonup> (Eval) "(fn '_ => raise Match)"
123
124
125subsubsection \<open>Equality and order on functions\<close>
126
127instantiation "fun" :: (enum, equal) equal
128begin
129
130definition
131  "HOL.equal f g \<longleftrightarrow> (\<forall>x \<in> set enum. f x = g x)"
132
133instance proof
134qed (simp_all add: equal_fun_def fun_eq_iff enum_UNIV)
135
136end
137
138lemma [code]:
139  "HOL.equal f g \<longleftrightarrow> enum_all (%x. f x = g x)"
140  by (auto simp add: equal fun_eq_iff)
141
142lemma [code nbe]:
143  "HOL.equal (f :: _ \<Rightarrow> _) f \<longleftrightarrow> True"
144  by (fact equal_refl)
145
146lemma order_fun [code]:
147  fixes f g :: "'a::enum \<Rightarrow> 'b::order"
148  shows "f \<le> g \<longleftrightarrow> enum_all (\<lambda>x. f x \<le> g x)"
149    and "f < g \<longleftrightarrow> f \<le> g \<and> enum_ex (\<lambda>x. f x \<noteq> g x)"
150  by (simp_all add: fun_eq_iff le_fun_def order_less_le)
151
152
153subsubsection \<open>Operations on relations\<close>
154
155lemma [code]:
156  "Id = image (\<lambda>x. (x, x)) (set Enum.enum)"
157  by (auto intro: imageI in_enum)
158
159lemma tranclp_unfold [code]:
160  "tranclp r a b \<longleftrightarrow> (a, b) \<in> trancl {(x, y). r x y}"
161  by (simp add: trancl_def)
162
163lemma rtranclp_rtrancl_eq [code]:
164  "rtranclp r x y \<longleftrightarrow> (x, y) \<in> rtrancl {(x, y). r x y}"
165  by (simp add: rtrancl_def)
166
167lemma 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
171lemma 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)
174
175lemma mlex_eq [code]:
176  "f <*mlex*> R = {(x, y). f x < f y \<or> (f x \<le> f y \<and> (x, y) \<in> R)}"
177  by (auto simp add: mlex_prod_def)
178
179
180subsubsection \<open>Bounded accessible part\<close>
181
182primrec bacc :: "('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> 'a set" 
183where
184  "bacc r 0 = {x. \<forall> y. (y, x) \<notin> r}"
185| "bacc r (Suc n) = (bacc r n \<union> {x. \<forall>y. (y, x) \<in> r \<longrightarrow> y \<in> bacc r n})"
186
187lemma bacc_subseteq_acc:
188  "bacc r n \<subseteq> Wellfounded.acc r"
189  by (induct n) (auto intro: acc.intros)
190
191lemma bacc_mono:
192  "n \<le> m \<Longrightarrow> bacc r n \<subseteq> bacc r m"
193  by (induct rule: dec_induct) auto
194  
195lemma bacc_upper_bound:
196  "bacc (r :: ('a \<times> 'a) set)  (card (UNIV :: 'a::finite set)) = (\<Union>n. bacc r n)"
197proof -
198  have "mono (bacc r)" unfolding mono_def by (simp add: bacc_mono)
199  moreover have "\<forall>n. bacc r n = bacc r (Suc n) \<longrightarrow> bacc r (Suc n) = bacc r (Suc (Suc n))" by auto
200  moreover have "finite (range (bacc r))" by auto
201  ultimately show ?thesis
202   by (intro finite_mono_strict_prefix_implies_finite_fixpoint)
203     (auto intro: finite_mono_remains_stable_implies_strict_prefix)
204qed
205
206lemma acc_subseteq_bacc:
207  assumes "finite r"
208  shows "Wellfounded.acc r \<subseteq> (\<Union>n. bacc r n)"
209proof
210  fix x
211  assume "x \<in> Wellfounded.acc r"
212  then have "\<exists>n. x \<in> bacc r n"
213  proof (induct x arbitrary: rule: acc.induct)
214    case (accI x)
215    then have "\<forall>y. \<exists> n. (y, x) \<in> r \<longrightarrow> y \<in> bacc r n" by simp
216    from choice[OF this] obtain n where n: "\<forall>y. (y, x) \<in> r \<longrightarrow> y \<in> bacc r (n y)" ..
217    obtain n where "\<And>y. (y, x) \<in> r \<Longrightarrow> y \<in> bacc r n"
218    proof
219      fix y assume y: "(y, x) \<in> r"
220      with n have "y \<in> bacc r (n y)" by auto
221      moreover have "n y <= Max ((\<lambda>(y, x). n y) ` r)"
222        using y \<open>finite r\<close> by (auto intro!: Max_ge)
223      note bacc_mono[OF this, of r]
224      ultimately show "y \<in> bacc r (Max ((\<lambda>(y, x). n y) ` r))" by auto
225    qed
226    then show ?case
227      by (auto simp add: Let_def intro!: exI[of _ "Suc n"])
228  qed
229  then show "x \<in> (\<Union>n. bacc r n)" by auto
230qed
231
232lemma acc_bacc_eq:
233  fixes A :: "('a :: finite \<times> 'a) set"
234  assumes "finite A"
235  shows "Wellfounded.acc A = bacc A (card (UNIV :: 'a set))"
236  using assms by (metis acc_subseteq_bacc bacc_subseteq_acc bacc_upper_bound order_eq_iff)
237
238lemma [code]:
239  fixes xs :: "('a::finite \<times> 'a) list"
240  shows "Wellfounded.acc (set xs) = bacc (set xs) (card_UNIV TYPE('a))"
241  by (simp add: card_UNIV_def acc_bacc_eq)
242
243
244subsection \<open>Default instances for @{class enum}\<close>
245
246lemma map_of_zip_enum_is_Some:
247  assumes "length ys = length (enum :: 'a::enum list)"
248  shows "\<exists>y. map_of (zip (enum :: 'a::enum list) ys) x = Some y"
249proof -
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)"
252    by (auto intro!: map_of_zip_is_Some)
253  then show ?thesis using enum_UNIV by auto
254qed
255
256lemma map_of_zip_enum_inject:
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)"
261  shows "xs = ys"
262proof -
263  have "map_of (zip (enum :: 'a list) xs) = map_of (zip (enum :: 'a list) ys)"
264  proof
265    fix x :: 'a
266    from length map_of_zip_enum_is_Some obtain y1 y2
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
269    moreover from map_of
270      have "the (map_of (zip (enum :: 'a::enum list) xs) x) = the (map_of (zip (enum :: 'a::enum list) ys) x)"
271      by (auto dest: fun_cong)
272    ultimately show "map_of (zip (enum :: 'a::enum list) xs) x = map_of (zip (enum :: 'a::enum list) ys) x"
273      by simp
274  qed
275  with length enum_distinct show "xs = ys" by (rule map_of_zip_inject)
276qed
277
278definition all_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
279where
280  "all_n_lists P n \<longleftrightarrow> (\<forall>xs \<in> set (List.n_lists n enum). P xs)"
281
282lemma [code]:
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)))"
284  unfolding all_n_lists_def enum_all
285  by (cases n) (auto simp add: enum_UNIV)
286
287definition ex_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
288where
289  "ex_n_lists P n \<longleftrightarrow> (\<exists>xs \<in> set (List.n_lists n enum). P xs)"
290
291lemma [code]:
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)))"
293  unfolding ex_n_lists_def enum_ex
294  by (cases n) (auto simp add: enum_UNIV)
295
296instantiation "fun" :: (enum, enum) enum
297begin
298
299definition
300  "enum = map (\<lambda>ys. the \<circ> map_of (zip (enum::'a list) ys)) (List.n_lists (length (enum::'a::enum list)) enum)"
301
302definition
303  "enum_all P = all_n_lists (\<lambda>bs. P (the \<circ> map_of (zip enum bs))) (length (enum :: 'a list))"
304
305definition
306  "enum_ex P = ex_n_lists (\<lambda>bs. P (the \<circ> map_of (zip enum bs))) (length (enum :: 'a list))"
307
308instance proof
309  show "UNIV = set (enum :: ('a \<Rightarrow> 'b) list)"
310  proof (rule UNIV_eq_I)
311    fix f :: "'a \<Rightarrow> 'b"
312    have "f = the \<circ> map_of (zip (enum :: 'a::enum list) (map f enum))"
313      by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum)
314    then show "f \<in> set enum"
315      by (auto simp add: enum_fun_def set_n_lists intro: in_enum)
316  qed
317next
318  from map_of_zip_enum_inject
319  show "distinct (enum :: ('a \<Rightarrow> 'b) list)"
320    by (auto intro!: inj_onI simp add: enum_fun_def
321      distinct_map distinct_n_lists enum_distinct set_n_lists)
322next
323  fix P
324  show "enum_all (P :: ('a \<Rightarrow> 'b) \<Rightarrow> bool) = Ball UNIV P"
325  proof
326    assume "enum_all P"
327    show "Ball UNIV P"
328    proof
329      fix f :: "'a \<Rightarrow> 'b"
330      have f: "f = the \<circ> map_of (zip (enum :: 'a::enum list) (map f enum))"
331        by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum)
332      from \<open>enum_all P\<close> have "P (the \<circ> map_of (zip enum (map f enum)))"
333        unfolding enum_all_fun_def all_n_lists_def
334        apply (simp add: set_n_lists)
335        apply (erule_tac x="map f enum" in allE)
336        apply (auto intro!: in_enum)
337        done
338      from this f show "P f" by auto
339    qed
340  next
341    assume "Ball UNIV P"
342    from this show "enum_all P"
343      unfolding enum_all_fun_def all_n_lists_def by auto
344  qed
345next
346  fix P
347  show "enum_ex (P :: ('a \<Rightarrow> 'b) \<Rightarrow> bool) = Bex UNIV P"
348  proof
349    assume "enum_ex P"
350    from this show "Bex UNIV P"
351      unfolding enum_ex_fun_def ex_n_lists_def by auto
352  next
353    assume "Bex UNIV P"
354    from this obtain f where "P f" ..
355    have f: "f = the \<circ> map_of (zip (enum :: 'a::enum list) (map f enum))"
356      by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) 
357    from \<open>P f\<close> this have "P (the \<circ> map_of (zip (enum :: 'a::enum list) (map f enum)))"
358      by auto
359    from  this show "enum_ex P"
360      unfolding enum_ex_fun_def ex_n_lists_def
361      apply (auto simp add: set_n_lists)
362      apply (rule_tac x="map f enum" in exI)
363      apply (auto intro!: in_enum)
364      done
365  qed
366qed
367
368end
369
370lemma enum_fun_code [code]: "enum = (let enum_a = (enum :: 'a::{enum, equal} list)
371  in map (\<lambda>ys. the \<circ> map_of (zip enum_a ys)) (List.n_lists (length enum_a) enum))"
372  by (simp add: enum_fun_def Let_def)
373
374lemma enum_all_fun_code [code]:
375  "enum_all P = (let enum_a = (enum :: 'a::{enum, equal} list)
376   in all_n_lists (\<lambda>bs. P (the \<circ> map_of (zip enum_a bs))) (length enum_a))"
377  by (simp only: enum_all_fun_def Let_def)
378
379lemma enum_ex_fun_code [code]:
380  "enum_ex P = (let enum_a = (enum :: 'a::{enum, equal} list)
381   in ex_n_lists (\<lambda>bs. P (the \<circ> map_of (zip enum_a bs))) (length enum_a))"
382  by (simp only: enum_ex_fun_def Let_def)
383
384instantiation set :: (enum) enum
385begin
386
387definition
388  "enum = map set (subseqs enum)"
389
390definition
391  "enum_all P \<longleftrightarrow> (\<forall>A\<in>set enum. P (A::'a set))"
392
393definition
394  "enum_ex P \<longleftrightarrow> (\<exists>A\<in>set enum. P (A::'a set))"
395
396instance proof
397qed (simp_all add: enum_set_def enum_all_set_def enum_ex_set_def subseqs_powset distinct_set_subseqs
398  enum_distinct enum_UNIV)
399
400end
401
402instantiation unit :: enum
403begin
404
405definition
406  "enum = [()]"
407
408definition
409  "enum_all P = P ()"
410
411definition
412  "enum_ex P = P ()"
413
414instance proof
415qed (auto simp add: enum_unit_def enum_all_unit_def enum_ex_unit_def)
416
417end
418
419instantiation bool :: enum
420begin
421
422definition
423  "enum = [False, True]"
424
425definition
426  "enum_all P \<longleftrightarrow> P False \<and> P True"
427
428definition
429  "enum_ex P \<longleftrightarrow> P False \<or> P True"
430
431instance proof
432qed (simp_all only: enum_bool_def enum_all_bool_def enum_ex_bool_def UNIV_bool, simp_all)
433
434end
435
436instantiation prod :: (enum, enum) enum
437begin
438
439definition
440  "enum = List.product enum enum"
441
442definition
443  "enum_all P = enum_all (%x. enum_all (%y. P (x, y)))"
444
445definition
446  "enum_ex P = enum_ex (%x. enum_ex (%y. P (x, y)))"
447
448 
449instance
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)
453
454end
455
456instantiation sum :: (enum, enum) enum
457begin
458
459definition
460  "enum = map Inl enum @ map Inr enum"
461
462definition
463  "enum_all P \<longleftrightarrow> enum_all (\<lambda>x. P (Inl x)) \<and> enum_all (\<lambda>x. P (Inr x))"
464
465definition
466  "enum_ex P \<longleftrightarrow> enum_ex (\<lambda>x. P (Inl x)) \<or> enum_ex (\<lambda>x. P (Inr x))"
467
468instance proof
469qed (simp_all only: enum_sum_def enum_all_sum_def enum_ex_sum_def UNIV_sum,
470  auto simp add: enum_UNIV distinct_map enum_distinct)
471
472end
473
474instantiation option :: (enum) enum
475begin
476
477definition
478  "enum = None # map Some enum"
479
480definition
481  "enum_all P \<longleftrightarrow> P None \<and> enum_all (\<lambda>x. P (Some x))"
482
483definition
484  "enum_ex P \<longleftrightarrow> P None \<or> enum_ex (\<lambda>x. P (Some x))"
485
486instance proof
487qed (simp_all only: enum_option_def enum_all_option_def enum_ex_option_def UNIV_option_conv,
488  auto simp add: distinct_map enum_UNIV enum_distinct)
489
490end
491
492
493subsection \<open>Small finite types\<close>
494
495text \<open>We define small finite types for use in Quickcheck\<close>
496
497datatype (plugins only: code "quickcheck" extraction) finite_1 =
498  a\<^sub>1
499
500notation (output) a\<^sub>1  ("a\<^sub>1")
501
502lemma UNIV_finite_1:
503  "UNIV = {a\<^sub>1}"
504  by (auto intro: finite_1.exhaust)
505
506instantiation finite_1 :: enum
507begin
508
509definition
510  "enum = [a\<^sub>1]"
511
512definition
513  "enum_all P = P a\<^sub>1"
514
515definition
516  "enum_ex P = P a\<^sub>1"
517
518instance proof
519qed (simp_all only: enum_finite_1_def enum_all_finite_1_def enum_ex_finite_1_def UNIV_finite_1, simp_all)
520
521end
522
523instantiation finite_1 :: linorder
524begin
525
526definition less_finite_1 :: "finite_1 \<Rightarrow> finite_1 \<Rightarrow> bool"
527where
528  "x < (y :: finite_1) \<longleftrightarrow> False"
529
530definition less_eq_finite_1 :: "finite_1 \<Rightarrow> finite_1 \<Rightarrow> bool"
531where
532  "x \<le> (y :: finite_1) \<longleftrightarrow> True"
533
534instance
535apply (intro_classes)
536apply (auto simp add: less_finite_1_def less_eq_finite_1_def)
537apply (metis finite_1.exhaust)
538done
539
540end
541
542instance finite_1 :: "{dense_linorder, wellorder}"
543by intro_classes (simp_all add: less_finite_1_def)
544
545instantiation finite_1 :: complete_lattice
546begin
547
548definition [simp]: "Inf = (\<lambda>_. a\<^sub>1)"
549definition [simp]: "Sup = (\<lambda>_. a\<^sub>1)"
550definition [simp]: "bot = a\<^sub>1"
551definition [simp]: "top = a\<^sub>1"
552definition [simp]: "inf = (\<lambda>_ _. a\<^sub>1)"
553definition [simp]: "sup = (\<lambda>_ _. a\<^sub>1)"
554
555instance by intro_classes(simp_all add: less_eq_finite_1_def)
556end
557
558instance finite_1 :: complete_distrib_lattice
559  by standard simp_all
560
561instance finite_1 :: complete_linorder ..
562
563lemma finite_1_eq: "x = a\<^sub>1"
564by(cases x) simp
565
566simproc_setup finite_1_eq ("x::finite_1") = \<open>
567  fn _ => fn _ => fn ct =>
568    (case Thm.term_of ct of
569      Const (@{const_name a\<^sub>1}, _) => NONE
570    | _ => SOME (mk_meta_eq @{thm finite_1_eq}))
571\<close>
572
573instantiation finite_1 :: complete_boolean_algebra
574begin
575definition [simp]: "(-) = (\<lambda>_ _. a\<^sub>1)"
576definition [simp]: "uminus = (\<lambda>_. a\<^sub>1)"
577instance by intro_classes simp_all
578end
579
580instantiation finite_1 :: 
581  "{linordered_ring_strict, linordered_comm_semiring_strict, ordered_comm_ring,
582    ordered_cancel_comm_monoid_diff, comm_monoid_mult, ordered_ring_abs,
583    one, modulo, sgn, inverse}"
584begin
585definition [simp]: "Groups.zero = a\<^sub>1"
586definition [simp]: "Groups.one = a\<^sub>1"
587definition [simp]: "(+) = (\<lambda>_ _. a\<^sub>1)"
588definition [simp]: "( * ) = (\<lambda>_ _. a\<^sub>1)"
589definition [simp]: "(mod) = (\<lambda>_ _. a\<^sub>1)" 
590definition [simp]: "abs = (\<lambda>_. a\<^sub>1)"
591definition [simp]: "sgn = (\<lambda>_. a\<^sub>1)"
592definition [simp]: "inverse = (\<lambda>_. a\<^sub>1)"
593definition [simp]: "divide = (\<lambda>_ _. a\<^sub>1)"
594
595instance by intro_classes(simp_all add: less_finite_1_def)
596end
597
598declare [[simproc del: finite_1_eq]]
599hide_const (open) a\<^sub>1
600
601datatype (plugins only: code "quickcheck" extraction) finite_2 =
602  a\<^sub>1 | a\<^sub>2
603
604notation (output) a\<^sub>1  ("a\<^sub>1")
605notation (output) a\<^sub>2  ("a\<^sub>2")
606
607lemma UNIV_finite_2:
608  "UNIV = {a\<^sub>1, a\<^sub>2}"
609  by (auto intro: finite_2.exhaust)
610
611instantiation finite_2 :: enum
612begin
613
614definition
615  "enum = [a\<^sub>1, a\<^sub>2]"
616
617definition
618  "enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2"
619
620definition
621  "enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2"
622
623instance proof
624qed (simp_all only: enum_finite_2_def enum_all_finite_2_def enum_ex_finite_2_def UNIV_finite_2, simp_all)
625
626end
627
628instantiation finite_2 :: linorder
629begin
630
631definition less_finite_2 :: "finite_2 \<Rightarrow> finite_2 \<Rightarrow> bool"
632where
633  "x < y \<longleftrightarrow> x = a\<^sub>1 \<and> y = a\<^sub>2"
634
635definition less_eq_finite_2 :: "finite_2 \<Rightarrow> finite_2 \<Rightarrow> bool"
636where
637  "x \<le> y \<longleftrightarrow> x = y \<or> x < (y :: finite_2)"
638
639instance
640apply (intro_classes)
641apply (auto simp add: less_finite_2_def less_eq_finite_2_def)
642apply (metis finite_2.nchotomy)+
643done
644
645end
646
647instance finite_2 :: wellorder
648by(rule wf_wellorderI)(simp add: less_finite_2_def, intro_classes)
649
650instantiation finite_2 :: complete_lattice
651begin
652
653definition "\<Sqinter>A = (if a\<^sub>1 \<in> A then a\<^sub>1 else a\<^sub>2)"
654definition "\<Squnion>A = (if a\<^sub>2 \<in> A then a\<^sub>2 else a\<^sub>1)"
655definition [simp]: "bot = a\<^sub>1"
656definition [simp]: "top = a\<^sub>2"
657definition "x \<sqinter> y = (if x = a\<^sub>1 \<or> y = a\<^sub>1 then a\<^sub>1 else a\<^sub>2)"
658definition "x \<squnion> y = (if x = a\<^sub>2 \<or> y = a\<^sub>2 then a\<^sub>2 else a\<^sub>1)"
659
660lemma neq_finite_2_a\<^sub>1_iff [simp]: "x \<noteq> a\<^sub>1 \<longleftrightarrow> x = a\<^sub>2"
661by(cases x) simp_all
662
663lemma neq_finite_2_a\<^sub>1_iff' [simp]: "a\<^sub>1 \<noteq> x \<longleftrightarrow> x = a\<^sub>2"
664by(cases x) simp_all
665
666lemma neq_finite_2_a\<^sub>2_iff [simp]: "x \<noteq> a\<^sub>2 \<longleftrightarrow> x = a\<^sub>1"
667by(cases x) simp_all
668
669lemma neq_finite_2_a\<^sub>2_iff' [simp]: "a\<^sub>2 \<noteq> x \<longleftrightarrow> x = a\<^sub>1"
670by(cases x) simp_all
671
672instance
673proof
674  fix x :: finite_2 and A
675  assume "x \<in> A"
676  then show "\<Sqinter>A \<le> x" "x \<le> \<Squnion>A"
677    by(case_tac [!] x)(auto simp add: less_eq_finite_2_def less_finite_2_def Inf_finite_2_def Sup_finite_2_def)
678qed(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)
679end
680
681instance finite_2 :: complete_linorder ..
682
683instance finite_2 :: complete_distrib_lattice ..
684
685instantiation finite_2 :: "{field, idom_abs_sgn, idom_modulo}" begin
686definition [simp]: "0 = a\<^sub>1"
687definition [simp]: "1 = a\<^sub>2"
688definition "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)"
689definition "uminus = (\<lambda>x :: finite_2. x)"
690definition "(-) = ((+) :: finite_2 \<Rightarrow> _)"
691definition "x * y = (case (x, y) of (a\<^sub>2, a\<^sub>2) \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> a\<^sub>1)"
692definition "inverse = (\<lambda>x :: finite_2. x)"
693definition "divide = (( * ) :: finite_2 \<Rightarrow> _)"
694definition "x mod y = (case (x, y) of (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> a\<^sub>1)"
695definition "abs = (\<lambda>x :: finite_2. x)"
696definition "sgn = (\<lambda>x :: finite_2. x)"
697instance
698  by standard
699    (simp_all add: plus_finite_2_def uminus_finite_2_def minus_finite_2_def
700      times_finite_2_def
701      inverse_finite_2_def divide_finite_2_def modulo_finite_2_def
702      abs_finite_2_def sgn_finite_2_def
703      split: finite_2.splits)
704end
705
706lemma two_finite_2 [simp]:
707  "2 = a\<^sub>1"
708  by (simp add: numeral.simps plus_finite_2_def)
709
710lemma dvd_finite_2_unfold:
711  "x dvd y \<longleftrightarrow> x = a\<^sub>2 \<or> y = a\<^sub>1"
712  by (auto simp add: dvd_def times_finite_2_def split: finite_2.splits)
713
714instantiation finite_2 :: "{normalization_semidom, unique_euclidean_semiring}" begin
715definition [simp]: "normalize = (id :: finite_2 \<Rightarrow> _)"
716definition [simp]: "unit_factor = (id :: finite_2 \<Rightarrow> _)"
717definition [simp]: "euclidean_size x = (case x of a\<^sub>1 \<Rightarrow> 0 | a\<^sub>2 \<Rightarrow> 1)"
718definition [simp]: "division_segment (x :: finite_2) = 1"
719instance
720  by standard
721    (auto simp add: divide_finite_2_def times_finite_2_def dvd_finite_2_unfold
722    split: finite_2.splits)
723end
724
725 
726hide_const (open) a\<^sub>1 a\<^sub>2
727
728datatype (plugins only: code "quickcheck" extraction) finite_3 =
729  a\<^sub>1 | a\<^sub>2 | a\<^sub>3
730
731notation (output) a\<^sub>1  ("a\<^sub>1")
732notation (output) a\<^sub>2  ("a\<^sub>2")
733notation (output) a\<^sub>3  ("a\<^sub>3")
734
735lemma UNIV_finite_3:
736  "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3}"
737  by (auto intro: finite_3.exhaust)
738
739instantiation finite_3 :: enum
740begin
741
742definition
743  "enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3]"
744
745definition
746  "enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2 \<and> P a\<^sub>3"
747
748definition
749  "enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2 \<or> P a\<^sub>3"
750
751instance proof
752qed (simp_all only: enum_finite_3_def enum_all_finite_3_def enum_ex_finite_3_def UNIV_finite_3, simp_all)
753
754end
755
756lemma finite_3_not_eq_unfold:
757  "x \<noteq> a\<^sub>1 \<longleftrightarrow> x \<in> {a\<^sub>2, a\<^sub>3}"
758  "x \<noteq> a\<^sub>2 \<longleftrightarrow> x \<in> {a\<^sub>1, a\<^sub>3}"
759  "x \<noteq> a\<^sub>3 \<longleftrightarrow> x \<in> {a\<^sub>1, a\<^sub>2}"
760  by (cases x; simp)+
761
762instantiation finite_3 :: linorder
763begin
764
765definition less_finite_3 :: "finite_3 \<Rightarrow> finite_3 \<Rightarrow> bool"
766where
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)"
768
769definition less_eq_finite_3 :: "finite_3 \<Rightarrow> finite_3 \<Rightarrow> bool"
770where
771  "x \<le> y \<longleftrightarrow> x = y \<or> x < (y :: finite_3)"
772
773instance proof (intro_classes)
774qed (auto simp add: less_finite_3_def less_eq_finite_3_def split: finite_3.split_asm)
775
776end
777
778instance finite_3 :: wellorder
779proof(rule wf_wellorderI)
780  have "inv_image less_than (case_finite_3 0 1 2) = {(x, y). x < y}"
781    by(auto simp add: less_finite_3_def split: finite_3.splits)
782  from this[symmetric] show "wf \<dots>" by simp
783qed intro_classes
784
785class finite_lattice = finite +  lattice + Inf + Sup  + bot + top +
786  assumes Inf_finite_empty: "Inf {} = Sup UNIV"
787  assumes Inf_finite_insert: "Inf (insert a A) = a \<sqinter> Inf A"
788  assumes Sup_finite_empty: "Sup {} = Inf UNIV"
789  assumes Sup_finite_insert: "Sup (insert a A) = a \<squnion> Sup A"
790  assumes bot_finite_def: "bot = Inf UNIV"
791  assumes top_finite_def: "top = Sup UNIV"
792begin
793
794subclass complete_lattice
795proof
796  fix x A
797  show "x \<in> A \<Longrightarrow> \<Sqinter>A \<le> x"
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)
799  show "x \<in> A \<Longrightarrow> x \<le> \<Squnion>A"
800    by (metis Set.set_insert insert_absorb2 local.Sup_finite_insert local.sup.absorb_iff2)
801next
802  fix A z
803  have "\<Squnion> UNIV = z \<squnion> \<Squnion>UNIV"
804    by (subst Sup_finite_insert [symmetric], simp add: insert_UNIV)
805  from this have [simp]: "z \<le> \<Squnion>UNIV"
806    using local.le_iff_sup by auto
807  have "(\<forall> x. x \<in> A \<longrightarrow> z \<le> x) \<longrightarrow> z \<le> \<Sqinter>A"
808    apply (rule finite_induct [of A "\<lambda> A . (\<forall> x. x \<in> A \<longrightarrow> z \<le> x) \<longrightarrow> z \<le> \<Sqinter>A"])
809    by (simp_all add: Inf_finite_empty Inf_finite_insert)
810  from this show "(\<And>x. x \<in> A \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> \<Sqinter>A"
811    by simp
812
813  have "\<Sqinter> UNIV = z \<sqinter> \<Sqinter>UNIV"
814    by (subst Inf_finite_insert [symmetric], simp add: insert_UNIV)
815  from this have [simp]: "\<Sqinter>UNIV \<le> z"
816    by (simp add: local.inf.absorb_iff2)
817  have "(\<forall> x. x \<in> A \<longrightarrow> x \<le> z) \<longrightarrow> \<Squnion>A \<le> z"
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)
819  from this show " (\<And>x. x \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> \<Squnion>A \<le> z"
820    by blast
821next
822  show "\<Sqinter>{} = \<top>"
823    by (simp add: Inf_finite_empty top_finite_def)
824  show " \<Squnion>{} = \<bottom>"
825    by (simp add: Sup_finite_empty bot_finite_def)
826qed
827end
828
829class finite_distrib_lattice = finite_lattice + distrib_lattice 
830begin
831lemma finite_inf_Sup: "a \<sqinter> (Sup A) = Sup {a \<sqinter> b | b . b \<in> A}"
832proof (rule finite_induct [of A "\<lambda> A . a \<sqinter> (Sup A) = Sup {a \<sqinter> b | b . b \<in> A}"], simp_all)
833  fix x::"'a"
834  fix F
835  assume "x \<notin> F"
836  assume [simp]: "a \<sqinter> \<Squnion>F = \<Squnion>{a \<sqinter> b |b. b \<in> F}"
837  have [simp]: " insert (a \<sqinter> x) {a \<sqinter> b |b. b \<in> F} = {a \<sqinter> b |b. b = x \<or> b \<in> F}"
838    by blast
839  have "a \<sqinter> (x \<squnion> \<Squnion>F) = a \<sqinter> x \<squnion> a \<sqinter> \<Squnion>F"
840    by (simp add: inf_sup_distrib1)
841  also have "... = a \<sqinter> x \<squnion> \<Squnion>{a \<sqinter> b |b. b \<in> F}"
842    by simp
843  also have "... = \<Squnion>{a \<sqinter> b |b. b = x \<or> b \<in> F}"
844    by (unfold Sup_insert[THEN sym], simp)
845  finally show "a \<sqinter> (x \<squnion> \<Squnion>F) = \<Squnion>{a \<sqinter> b |b. b = x \<or> b \<in> F}"
846    by simp
847qed
848
849lemma finite_Inf_Sup: "INFIMUM A Sup \<le> SUPREMUM {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y} Inf"
850proof (rule  finite_induct [of A "\<lambda> A .INFIMUM A Sup \<le> SUPREMUM {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y} Inf"], simp_all add: finite_UnionD)
851  fix x::"'a set"
852  fix F
853  assume "x \<notin> F"
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)}"
855    by auto
856  define fa where "fa = (\<lambda> (b::'a) f Y . (if Y = x then b else f Y))"
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)"
858    by (auto simp add: fa_def)
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)}"
860    by blast
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> SUPREMUM {insert (f x) (f ` F) |f. f x \<in> x \<and> (\<forall>Y\<in>F. f Y \<in> Y)} Inf"
862    using B apply (rule SUP_upper2, simp_all)
863    by (simp_all add: INF_greatest Inf_lower inf.coboundedI2 fa_def)
864
865  assume "INFIMUM F Sup \<le> SUPREMUM {f ` F |f. \<forall>Y\<in>F. f Y \<in> Y} Inf"
866
867  from this have "\<Squnion>x \<sqinter> INFIMUM F Sup \<le> \<Squnion>x \<sqinter> SUPREMUM {f ` F |f. \<forall>Y\<in>F. f Y \<in> Y} Inf"
868    apply simp
869    using inf.coboundedI2 by blast
870  also have "... = Sup {\<Squnion>x \<sqinter> (Inf (f ` F)) |f  .  (\<forall>Y\<in>F. f Y \<in> Y)}"
871    by (simp add: finite_inf_Sup)
872
873  also have "... = Sup {Sup {Inf (f ` F) \<sqinter> b | b . b \<in> x} |f  .  (\<forall>Y\<in>F. f Y \<in> Y)}"
874    apply (subst inf_commute)
875    by (simp add: finite_inf_Sup)
876
877  also have "... \<le> SUPREMUM {insert (f x) (f ` F) |f. f x \<in> x \<and> (\<forall>Y\<in>F. f Y \<in> Y)} Inf"
878    apply (rule Sup_least, clarsimp)+
879    by (subst inf_commute, simp)
880
881  finally show "\<Squnion>x \<sqinter> INFIMUM F Sup \<le> SUPREMUM {insert (f x) (f ` F) |f. f x \<in> x \<and> (\<forall>Y\<in>F. f Y \<in> Y)} Inf"
882    by simp
883qed
884
885subclass complete_distrib_lattice
886  by (standard, rule finite_Inf_Sup)
887end
888
889instantiation finite_3 :: finite_lattice
890begin
891
892definition "\<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)"
893definition "\<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)"
894definition [simp]: "bot = a\<^sub>1"
895definition [simp]: "top = a\<^sub>3"
896definition [simp]: "inf = (min :: finite_3 \<Rightarrow> _)"
897definition [simp]: "sup = (max :: finite_3 \<Rightarrow> _)"
898
899instance
900proof
901qed (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)
902end
903
904instance finite_3 :: complete_lattice ..
905
906instance finite_3 :: finite_distrib_lattice
907proof 
908qed (auto simp add: min_def max_def)
909
910instance finite_3 :: complete_distrib_lattice ..
911
912instance finite_3 :: complete_linorder ..
913
914instantiation finite_3 :: "{field, idom_abs_sgn, idom_modulo}" begin
915definition [simp]: "0 = a\<^sub>1"
916definition [simp]: "1 = a\<^sub>2"
917definition
918  "x + y = (case (x, y) of
919     (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
920   | (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
921   | _ \<Rightarrow> a\<^sub>3)"
922definition "- 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)"
923definition "x - y = x + (- y :: finite_3)"
924definition "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)"
925definition "inverse = (\<lambda>x :: finite_3. x)" 
926definition "x div y = x * inverse (y :: finite_3)"
927definition "x mod y = (case y of a\<^sub>1 \<Rightarrow> x | _ \<Rightarrow> a\<^sub>1)"
928definition "abs = (\<lambda>x. case x of a\<^sub>3 \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> x)"
929definition "sgn = (\<lambda>x :: finite_3. x)"
930instance
931  by standard
932    (simp_all add: plus_finite_3_def uminus_finite_3_def minus_finite_3_def
933      times_finite_3_def
934      inverse_finite_3_def divide_finite_3_def modulo_finite_3_def
935      abs_finite_3_def sgn_finite_3_def
936      less_finite_3_def
937      split: finite_3.splits)
938end
939
940lemma two_finite_3 [simp]:
941  "2 = a\<^sub>3"
942  by (simp add: numeral.simps plus_finite_3_def)
943
944lemma dvd_finite_3_unfold:
945  "x dvd y \<longleftrightarrow> x = a\<^sub>2 \<or> x = a\<^sub>3 \<or> y = a\<^sub>1"
946  by (cases x) (auto simp add: dvd_def times_finite_3_def split: finite_3.splits)
947
948instantiation finite_3 :: "{normalization_semidom, unique_euclidean_semiring}" begin
949definition [simp]: "normalize x = (case x of a\<^sub>3 \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> x)"
950definition [simp]: "unit_factor = (id :: finite_3 \<Rightarrow> _)"
951definition [simp]: "euclidean_size x = (case x of a\<^sub>1 \<Rightarrow> 0 | _ \<Rightarrow> 1)"
952definition [simp]: "division_segment (x :: finite_3) = 1"
953instance proof
954  fix x :: finite_3
955  assume "x \<noteq> 0"
956  then show "is_unit (unit_factor x)"
957    by (cases x) (simp_all add: dvd_finite_3_unfold)
958qed (auto simp add: divide_finite_3_def times_finite_3_def
959  dvd_finite_3_unfold inverse_finite_3_def plus_finite_3_def
960  split: finite_3.splits)
961end
962
963hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3
964
965datatype (plugins only: code "quickcheck" extraction) finite_4 =
966  a\<^sub>1 | a\<^sub>2 | a\<^sub>3 | a\<^sub>4
967
968notation (output) a\<^sub>1  ("a\<^sub>1")
969notation (output) a\<^sub>2  ("a\<^sub>2")
970notation (output) a\<^sub>3  ("a\<^sub>3")
971notation (output) a\<^sub>4  ("a\<^sub>4")
972
973lemma UNIV_finite_4:
974  "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4}"
975  by (auto intro: finite_4.exhaust)
976
977instantiation finite_4 :: enum
978begin
979
980definition
981  "enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4]"
982
983definition
984  "enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2 \<and> P a\<^sub>3 \<and> P a\<^sub>4"
985
986definition
987  "enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2 \<or> P a\<^sub>3 \<or> P a\<^sub>4"
988
989instance proof
990qed (simp_all only: enum_finite_4_def enum_all_finite_4_def enum_ex_finite_4_def UNIV_finite_4, simp_all)
991
992end
993
994instantiation finite_4 :: finite_distrib_lattice begin
995
996text \<open>@{term a\<^sub>1} $<$ @{term a\<^sub>2},@{term a\<^sub>3} $<$ @{term a\<^sub>4},
997  but @{term a\<^sub>2} and @{term a\<^sub>3} are incomparable.\<close>
998
999definition
1000  "x < y \<longleftrightarrow> (case (x, y) of
1001     (a\<^sub>1, a\<^sub>1) \<Rightarrow> False | (a\<^sub>1, _) \<Rightarrow> True
1002   |  (a\<^sub>2, a\<^sub>4) \<Rightarrow> True
1003   |  (a\<^sub>3, a\<^sub>4) \<Rightarrow> True  | _ \<Rightarrow> False)"
1004
1005definition 
1006  "x \<le> y \<longleftrightarrow> (case (x, y) of
1007     (a\<^sub>1, _) \<Rightarrow> True
1008   | (a\<^sub>2, a\<^sub>2) \<Rightarrow> True | (a\<^sub>2, a\<^sub>4) \<Rightarrow> True
1009   | (a\<^sub>3, a\<^sub>3) \<Rightarrow> True | (a\<^sub>3, a\<^sub>4) \<Rightarrow> True
1010   | (a\<^sub>4, a\<^sub>4) \<Rightarrow> True | _ \<Rightarrow> False)"
1011
1012definition
1013  "\<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)"
1014definition
1015  "\<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)"
1016definition [simp]: "bot = a\<^sub>1"
1017definition [simp]: "top = a\<^sub>4"
1018definition
1019  "x \<sqinter> y = (case (x, y) of
1020     (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
1021   | (a\<^sub>2, _) \<Rightarrow> a\<^sub>2 | (_, a\<^sub>2) \<Rightarrow> a\<^sub>2
1022   | (a\<^sub>3, _) \<Rightarrow> a\<^sub>3 | (_, a\<^sub>3) \<Rightarrow> a\<^sub>3
1023   | _ \<Rightarrow> a\<^sub>4)"
1024definition
1025  "x \<squnion> y = (case (x, y) of
1026     (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
1027  | (a\<^sub>2, _) \<Rightarrow> a\<^sub>2 | (_, a\<^sub>2) \<Rightarrow> a\<^sub>2
1028  | (a\<^sub>3, _) \<Rightarrow> a\<^sub>3 | (_, a\<^sub>3) \<Rightarrow> a\<^sub>3
1029  | _ \<Rightarrow> a\<^sub>1)"
1030
1031instance proof
1032qed(auto simp add: less_finite_4_def less_eq_finite_4_def Inf_finite_4_def Sup_finite_4_def 
1033    inf_finite_4_def sup_finite_4_def split: finite_4.splits)
1034end
1035
1036instance finite_4 :: complete_lattice ..
1037
1038instance finite_4 :: complete_distrib_lattice ..
1039
1040instantiation finite_4 :: complete_boolean_algebra begin
1041definition "- 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)"
1042definition "x - y = x \<sqinter> - (y :: finite_4)"
1043instance
1044by intro_classes
1045  (simp_all add: inf_finite_4_def sup_finite_4_def uminus_finite_4_def minus_finite_4_def 
1046    split: finite_4.splits)
1047end
1048
1049hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4
1050
1051datatype (plugins only: code "quickcheck" extraction) finite_5 =
1052  a\<^sub>1 | a\<^sub>2 | a\<^sub>3 | a\<^sub>4 | a\<^sub>5
1053
1054notation (output) a\<^sub>1  ("a\<^sub>1")
1055notation (output) a\<^sub>2  ("a\<^sub>2")
1056notation (output) a\<^sub>3  ("a\<^sub>3")
1057notation (output) a\<^sub>4  ("a\<^sub>4")
1058notation (output) a\<^sub>5  ("a\<^sub>5")
1059
1060lemma UNIV_finite_5:
1061  "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4, a\<^sub>5}"
1062  by (auto intro: finite_5.exhaust)
1063
1064instantiation finite_5 :: enum
1065begin
1066
1067definition
1068  "enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4, a\<^sub>5]"
1069
1070definition
1071  "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"
1072
1073definition
1074  "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"
1075
1076instance proof
1077qed (simp_all only: enum_finite_5_def enum_all_finite_5_def enum_ex_finite_5_def UNIV_finite_5, simp_all)
1078
1079end
1080
1081instantiation finite_5 :: finite_lattice
1082begin
1083
1084text \<open>The non-distributive pentagon lattice $N_5$\<close>
1085
1086definition
1087  "x < y \<longleftrightarrow> (case (x, y) of
1088     (a\<^sub>1, a\<^sub>1) \<Rightarrow> False | (a\<^sub>1, _) \<Rightarrow> True
1089   | (a\<^sub>2, a\<^sub>3) \<Rightarrow> True  | (a\<^sub>2, a\<^sub>5) \<Rightarrow> True
1090   | (a\<^sub>3, a\<^sub>5) \<Rightarrow> True
1091   | (a\<^sub>4, a\<^sub>5) \<Rightarrow> True  | _ \<Rightarrow> False)"
1092
1093definition
1094  "x \<le> y \<longleftrightarrow> (case (x, y) of
1095     (a\<^sub>1, _) \<Rightarrow> True
1096   | (a\<^sub>2, a\<^sub>2) \<Rightarrow> True | (a\<^sub>2, a\<^sub>3) \<Rightarrow> True | (a\<^sub>2, a\<^sub>5) \<Rightarrow> True
1097   | (a\<^sub>3, a\<^sub>3) \<Rightarrow> True | (a\<^sub>3, a\<^sub>5) \<Rightarrow> True
1098   | (a\<^sub>4, a\<^sub>4) \<Rightarrow> True | (a\<^sub>4, a\<^sub>5) \<Rightarrow> True
1099   | (a\<^sub>5, a\<^sub>5) \<Rightarrow> True | _ \<Rightarrow> False)"
1100
1101definition
1102  "\<Sqinter>A = 
1103  (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
1104   else if a\<^sub>2 \<in> A then a\<^sub>2
1105   else if a\<^sub>3 \<in> A then a\<^sub>3
1106   else if a\<^sub>4 \<in> A then a\<^sub>4
1107   else a\<^sub>5)"
1108definition
1109  "\<Squnion>A = 
1110  (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
1111   else if a\<^sub>3 \<in> A then a\<^sub>3
1112   else if a\<^sub>2 \<in> A then a\<^sub>2
1113   else if a\<^sub>4 \<in> A then a\<^sub>4
1114   else a\<^sub>1)"
1115definition [simp]: "bot = a\<^sub>1"
1116definition [simp]: "top = a\<^sub>5"
1117definition
1118  "x \<sqinter> y = (case (x, y) of
1119     (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
1120   | (a\<^sub>2, _) \<Rightarrow> a\<^sub>2 | (_, a\<^sub>2) \<Rightarrow> a\<^sub>2
1121   | (a\<^sub>3, _) \<Rightarrow> a\<^sub>3 | (_, a\<^sub>3) \<Rightarrow> a\<^sub>3
1122   | (a\<^sub>4, _) \<Rightarrow> a\<^sub>4 | (_, a\<^sub>4) \<Rightarrow> a\<^sub>4
1123   | _ \<Rightarrow> a\<^sub>5)"
1124definition
1125  "x \<squnion> y = (case (x, y) of
1126     (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
1127   | (a\<^sub>3, _) \<Rightarrow> a\<^sub>3 | (_, a\<^sub>3) \<Rightarrow> a\<^sub>3
1128   | (a\<^sub>2, _) \<Rightarrow> a\<^sub>2 | (_, a\<^sub>2) \<Rightarrow> a\<^sub>2
1129   | (a\<^sub>4, _) \<Rightarrow> a\<^sub>4 | (_, a\<^sub>4) \<Rightarrow> a\<^sub>4
1130   | _ \<Rightarrow> a\<^sub>1)"
1131
1132instance 
1133proof
1134qed (auto simp add: less_eq_finite_5_def less_finite_5_def inf_finite_5_def sup_finite_5_def 
1135    Inf_finite_5_def Sup_finite_5_def split: finite_5.splits if_split_asm)
1136end
1137
1138
1139instance  finite_5 :: complete_lattice ..
1140
1141
1142hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4 a\<^sub>5
1143
1144
1145subsection \<open>Closing up\<close>
1146
1147hide_type (open) finite_1 finite_2 finite_3 finite_4 finite_5
1148hide_const (open) enum enum_all enum_ex all_n_lists ex_n_lists ntrancl
1149
1150end
1151