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