src/HOL/Enum.thy
author haftmann
Sun, 10 Nov 2013 15:05:06 +0100
changeset 55747 45a5523d4a63
parent 55600 c8cc5ab4a863
child 56232 cb892d835803
permissions -rw-r--r--
qualifed popular user space names
     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]:
   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]:
   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 "Wellfounded.acc (set xs) = bacc (set xs) (card_UNIV TYPE('a))"
   182   by (simp add: card_UNIV_def acc_bacc_eq)
   183 
   184 
   185 subsection {* Default instances for @{class enum} *}
   186 
   187 lemma map_of_zip_enum_is_Some:
   188   assumes "length ys = length (enum \<Colon> 'a\<Colon>enum list)"
   189   shows "\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y"
   190 proof -
   191   from assms have "x \<in> set (enum \<Colon> 'a\<Colon>enum list) \<longleftrightarrow>
   192     (\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y)"
   193     by (auto intro!: map_of_zip_is_Some)
   194   then show ?thesis using enum_UNIV by auto
   195 qed
   196 
   197 lemma map_of_zip_enum_inject:
   198   fixes xs ys :: "'b\<Colon>enum list"
   199   assumes length: "length xs = length (enum \<Colon> 'a\<Colon>enum list)"
   200       "length ys = length (enum \<Colon> 'a\<Colon>enum list)"
   201     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)"
   202   shows "xs = ys"
   203 proof -
   204   have "map_of (zip (enum \<Colon> 'a list) xs) = map_of (zip (enum \<Colon> 'a list) ys)"
   205   proof
   206     fix x :: 'a
   207     from length map_of_zip_enum_is_Some obtain y1 y2
   208       where "map_of (zip (enum \<Colon> 'a list) xs) x = Some y1"
   209         and "map_of (zip (enum \<Colon> 'a list) ys) x = Some y2" by blast
   210     moreover from map_of
   211       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)"
   212       by (auto dest: fun_cong)
   213     ultimately show "map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) x = map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x"
   214       by simp
   215   qed
   216   with length enum_distinct show "xs = ys" by (rule map_of_zip_inject)
   217 qed
   218 
   219 definition all_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
   220 where
   221   "all_n_lists P n \<longleftrightarrow> (\<forall>xs \<in> set (List.n_lists n enum). P xs)"
   222 
   223 lemma [code]:
   224   "all_n_lists P n \<longleftrightarrow> (if n = 0 then P [] else enum_all (%x. all_n_lists (%xs. P (x # xs)) (n - 1)))"
   225   unfolding all_n_lists_def enum_all
   226   by (cases n) (auto simp add: enum_UNIV)
   227 
   228 definition ex_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
   229 where
   230   "ex_n_lists P n \<longleftrightarrow> (\<exists>xs \<in> set (List.n_lists n enum). P xs)"
   231 
   232 lemma [code]:
   233   "ex_n_lists P n \<longleftrightarrow> (if n = 0 then P [] else enum_ex (%x. ex_n_lists (%xs. P (x # xs)) (n - 1)))"
   234   unfolding ex_n_lists_def enum_ex
   235   by (cases n) (auto simp add: enum_UNIV)
   236 
   237 instantiation "fun" :: (enum, enum) enum
   238 begin
   239 
   240 definition
   241   "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)"
   242 
   243 definition
   244   "enum_all P = all_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))"
   245 
   246 definition
   247   "enum_ex P = ex_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))"
   248 
   249 instance proof
   250   show "UNIV = set (enum \<Colon> ('a \<Rightarrow> 'b) list)"
   251   proof (rule UNIV_eq_I)
   252     fix f :: "'a \<Rightarrow> 'b"
   253     have "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))"
   254       by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum)
   255     then show "f \<in> set enum"
   256       by (auto simp add: enum_fun_def set_n_lists intro: in_enum)
   257   qed
   258 next
   259   from map_of_zip_enum_inject
   260   show "distinct (enum \<Colon> ('a \<Rightarrow> 'b) list)"
   261     by (auto intro!: inj_onI simp add: enum_fun_def
   262       distinct_map distinct_n_lists enum_distinct set_n_lists)
   263 next
   264   fix P
   265   show "enum_all (P :: ('a \<Rightarrow> 'b) \<Rightarrow> bool) = Ball UNIV P"
   266   proof
   267     assume "enum_all P"
   268     show "Ball UNIV P"
   269     proof
   270       fix f :: "'a \<Rightarrow> 'b"
   271       have f: "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))"
   272         by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum)
   273       from `enum_all P` have "P (the \<circ> map_of (zip enum (map f enum)))"
   274         unfolding enum_all_fun_def all_n_lists_def
   275         apply (simp add: set_n_lists)
   276         apply (erule_tac x="map f enum" in allE)
   277         apply (auto intro!: in_enum)
   278         done
   279       from this f show "P f" by auto
   280     qed
   281   next
   282     assume "Ball UNIV P"
   283     from this show "enum_all P"
   284       unfolding enum_all_fun_def all_n_lists_def by auto
   285   qed
   286 next
   287   fix P
   288   show "enum_ex (P :: ('a \<Rightarrow> 'b) \<Rightarrow> bool) = Bex UNIV P"
   289   proof
   290     assume "enum_ex P"
   291     from this show "Bex UNIV P"
   292       unfolding enum_ex_fun_def ex_n_lists_def by auto
   293   next
   294     assume "Bex UNIV P"
   295     from this obtain f where "P f" ..
   296     have f: "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))"
   297       by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) 
   298     from `P f` this have "P (the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum)))"
   299       by auto
   300     from  this show "enum_ex P"
   301       unfolding enum_ex_fun_def ex_n_lists_def
   302       apply (auto simp add: set_n_lists)
   303       apply (rule_tac x="map f enum" in exI)
   304       apply (auto intro!: in_enum)
   305       done
   306   qed
   307 qed
   308 
   309 end
   310 
   311 lemma enum_fun_code [code]: "enum = (let enum_a = (enum \<Colon> 'a\<Colon>{enum, equal} list)
   312   in map (\<lambda>ys. the o map_of (zip enum_a ys)) (List.n_lists (length enum_a) enum))"
   313   by (simp add: enum_fun_def Let_def)
   314 
   315 lemma enum_all_fun_code [code]:
   316   "enum_all P = (let enum_a = (enum :: 'a::{enum, equal} list)
   317    in all_n_lists (\<lambda>bs. P (the o map_of (zip enum_a bs))) (length enum_a))"
   318   by (simp only: enum_all_fun_def Let_def)
   319 
   320 lemma enum_ex_fun_code [code]:
   321   "enum_ex P = (let enum_a = (enum :: 'a::{enum, equal} list)
   322    in ex_n_lists (\<lambda>bs. P (the o map_of (zip enum_a bs))) (length enum_a))"
   323   by (simp only: enum_ex_fun_def Let_def)
   324 
   325 instantiation set :: (enum) enum
   326 begin
   327 
   328 definition
   329   "enum = map set (sublists enum)"
   330 
   331 definition
   332   "enum_all P \<longleftrightarrow> (\<forall>A\<in>set enum. P (A::'a set))"
   333 
   334 definition
   335   "enum_ex P \<longleftrightarrow> (\<exists>A\<in>set enum. P (A::'a set))"
   336 
   337 instance proof
   338 qed (simp_all add: enum_set_def enum_all_set_def enum_ex_set_def sublists_powset distinct_set_sublists
   339   enum_distinct enum_UNIV)
   340 
   341 end
   342 
   343 instantiation unit :: enum
   344 begin
   345 
   346 definition
   347   "enum = [()]"
   348 
   349 definition
   350   "enum_all P = P ()"
   351 
   352 definition
   353   "enum_ex P = P ()"
   354 
   355 instance proof
   356 qed (auto simp add: enum_unit_def enum_all_unit_def enum_ex_unit_def)
   357 
   358 end
   359 
   360 instantiation bool :: enum
   361 begin
   362 
   363 definition
   364   "enum = [False, True]"
   365 
   366 definition
   367   "enum_all P \<longleftrightarrow> P False \<and> P True"
   368 
   369 definition
   370   "enum_ex P \<longleftrightarrow> P False \<or> P True"
   371 
   372 instance proof
   373 qed (simp_all only: enum_bool_def enum_all_bool_def enum_ex_bool_def UNIV_bool, simp_all)
   374 
   375 end
   376 
   377 instantiation prod :: (enum, enum) enum
   378 begin
   379 
   380 definition
   381   "enum = List.product enum enum"
   382 
   383 definition
   384   "enum_all P = enum_all (%x. enum_all (%y. P (x, y)))"
   385 
   386 definition
   387   "enum_ex P = enum_ex (%x. enum_ex (%y. P (x, y)))"
   388 
   389  
   390 instance by default
   391   (simp_all add: enum_prod_def product_list_set distinct_product
   392     enum_UNIV enum_distinct enum_all_prod_def enum_ex_prod_def)
   393 
   394 end
   395 
   396 instantiation sum :: (enum, enum) enum
   397 begin
   398 
   399 definition
   400   "enum = map Inl enum @ map Inr enum"
   401 
   402 definition
   403   "enum_all P \<longleftrightarrow> enum_all (\<lambda>x. P (Inl x)) \<and> enum_all (\<lambda>x. P (Inr x))"
   404 
   405 definition
   406   "enum_ex P \<longleftrightarrow> enum_ex (\<lambda>x. P (Inl x)) \<or> enum_ex (\<lambda>x. P (Inr x))"
   407 
   408 instance proof
   409 qed (simp_all only: enum_sum_def enum_all_sum_def enum_ex_sum_def UNIV_sum,
   410   auto simp add: enum_UNIV distinct_map enum_distinct)
   411 
   412 end
   413 
   414 instantiation option :: (enum) enum
   415 begin
   416 
   417 definition
   418   "enum = None # map Some enum"
   419 
   420 definition
   421   "enum_all P \<longleftrightarrow> P None \<and> enum_all (\<lambda>x. P (Some x))"
   422 
   423 definition
   424   "enum_ex P \<longleftrightarrow> P None \<or> enum_ex (\<lambda>x. P (Some x))"
   425 
   426 instance proof
   427 qed (simp_all only: enum_option_def enum_all_option_def enum_ex_option_def UNIV_option_conv,
   428   auto simp add: distinct_map enum_UNIV enum_distinct)
   429 
   430 end
   431 
   432 
   433 subsection {* Small finite types *}
   434 
   435 text {* We define small finite types for the use in Quickcheck *}
   436 
   437 datatype finite_1 = a\<^sub>1
   438 
   439 notation (output) a\<^sub>1  ("a\<^sub>1")
   440 
   441 lemma UNIV_finite_1:
   442   "UNIV = {a\<^sub>1}"
   443   by (auto intro: finite_1.exhaust)
   444 
   445 instantiation finite_1 :: enum
   446 begin
   447 
   448 definition
   449   "enum = [a\<^sub>1]"
   450 
   451 definition
   452   "enum_all P = P a\<^sub>1"
   453 
   454 definition
   455   "enum_ex P = P a\<^sub>1"
   456 
   457 instance proof
   458 qed (simp_all only: enum_finite_1_def enum_all_finite_1_def enum_ex_finite_1_def UNIV_finite_1, simp_all)
   459 
   460 end
   461 
   462 instantiation finite_1 :: linorder
   463 begin
   464 
   465 definition less_finite_1 :: "finite_1 \<Rightarrow> finite_1 \<Rightarrow> bool"
   466 where
   467   "x < (y :: finite_1) \<longleftrightarrow> False"
   468 
   469 definition less_eq_finite_1 :: "finite_1 \<Rightarrow> finite_1 \<Rightarrow> bool"
   470 where
   471   "x \<le> (y :: finite_1) \<longleftrightarrow> True"
   472 
   473 instance
   474 apply (intro_classes)
   475 apply (auto simp add: less_finite_1_def less_eq_finite_1_def)
   476 apply (metis finite_1.exhaust)
   477 done
   478 
   479 end
   480 
   481 hide_const (open) a\<^sub>1
   482 
   483 datatype finite_2 = a\<^sub>1 | a\<^sub>2
   484 
   485 notation (output) a\<^sub>1  ("a\<^sub>1")
   486 notation (output) a\<^sub>2  ("a\<^sub>2")
   487 
   488 lemma UNIV_finite_2:
   489   "UNIV = {a\<^sub>1, a\<^sub>2}"
   490   by (auto intro: finite_2.exhaust)
   491 
   492 instantiation finite_2 :: enum
   493 begin
   494 
   495 definition
   496   "enum = [a\<^sub>1, a\<^sub>2]"
   497 
   498 definition
   499   "enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2"
   500 
   501 definition
   502   "enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2"
   503 
   504 instance proof
   505 qed (simp_all only: enum_finite_2_def enum_all_finite_2_def enum_ex_finite_2_def UNIV_finite_2, simp_all)
   506 
   507 end
   508 
   509 instantiation finite_2 :: linorder
   510 begin
   511 
   512 definition less_finite_2 :: "finite_2 \<Rightarrow> finite_2 \<Rightarrow> bool"
   513 where
   514   "x < y \<longleftrightarrow> x = a\<^sub>1 \<and> y = a\<^sub>2"
   515 
   516 definition less_eq_finite_2 :: "finite_2 \<Rightarrow> finite_2 \<Rightarrow> bool"
   517 where
   518   "x \<le> y \<longleftrightarrow> x = y \<or> x < (y :: finite_2)"
   519 
   520 instance
   521 apply (intro_classes)
   522 apply (auto simp add: less_finite_2_def less_eq_finite_2_def)
   523 apply (metis finite_2.nchotomy)+
   524 done
   525 
   526 end
   527 
   528 hide_const (open) a\<^sub>1 a\<^sub>2
   529 
   530 datatype finite_3 = a\<^sub>1 | a\<^sub>2 | a\<^sub>3
   531 
   532 notation (output) a\<^sub>1  ("a\<^sub>1")
   533 notation (output) a\<^sub>2  ("a\<^sub>2")
   534 notation (output) a\<^sub>3  ("a\<^sub>3")
   535 
   536 lemma UNIV_finite_3:
   537   "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3}"
   538   by (auto intro: finite_3.exhaust)
   539 
   540 instantiation finite_3 :: enum
   541 begin
   542 
   543 definition
   544   "enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3]"
   545 
   546 definition
   547   "enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2 \<and> P a\<^sub>3"
   548 
   549 definition
   550   "enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2 \<or> P a\<^sub>3"
   551 
   552 instance proof
   553 qed (simp_all only: enum_finite_3_def enum_all_finite_3_def enum_ex_finite_3_def UNIV_finite_3, simp_all)
   554 
   555 end
   556 
   557 instantiation finite_3 :: linorder
   558 begin
   559 
   560 definition less_finite_3 :: "finite_3 \<Rightarrow> finite_3 \<Rightarrow> bool"
   561 where
   562   "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)"
   563 
   564 definition less_eq_finite_3 :: "finite_3 \<Rightarrow> finite_3 \<Rightarrow> bool"
   565 where
   566   "x \<le> y \<longleftrightarrow> x = y \<or> x < (y :: finite_3)"
   567 
   568 instance proof (intro_classes)
   569 qed (auto simp add: less_finite_3_def less_eq_finite_3_def split: finite_3.split_asm)
   570 
   571 end
   572 
   573 hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3
   574 
   575 datatype finite_4 = a\<^sub>1 | a\<^sub>2 | a\<^sub>3 | a\<^sub>4
   576 
   577 notation (output) a\<^sub>1  ("a\<^sub>1")
   578 notation (output) a\<^sub>2  ("a\<^sub>2")
   579 notation (output) a\<^sub>3  ("a\<^sub>3")
   580 notation (output) a\<^sub>4  ("a\<^sub>4")
   581 
   582 lemma UNIV_finite_4:
   583   "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4}"
   584   by (auto intro: finite_4.exhaust)
   585 
   586 instantiation finite_4 :: enum
   587 begin
   588 
   589 definition
   590   "enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4]"
   591 
   592 definition
   593   "enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2 \<and> P a\<^sub>3 \<and> P a\<^sub>4"
   594 
   595 definition
   596   "enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2 \<or> P a\<^sub>3 \<or> P a\<^sub>4"
   597 
   598 instance proof
   599 qed (simp_all only: enum_finite_4_def enum_all_finite_4_def enum_ex_finite_4_def UNIV_finite_4, simp_all)
   600 
   601 end
   602 
   603 hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4
   604 
   605 
   606 datatype finite_5 = a\<^sub>1 | a\<^sub>2 | a\<^sub>3 | a\<^sub>4 | a\<^sub>5
   607 
   608 notation (output) a\<^sub>1  ("a\<^sub>1")
   609 notation (output) a\<^sub>2  ("a\<^sub>2")
   610 notation (output) a\<^sub>3  ("a\<^sub>3")
   611 notation (output) a\<^sub>4  ("a\<^sub>4")
   612 notation (output) a\<^sub>5  ("a\<^sub>5")
   613 
   614 lemma UNIV_finite_5:
   615   "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4, a\<^sub>5}"
   616   by (auto intro: finite_5.exhaust)
   617 
   618 instantiation finite_5 :: enum
   619 begin
   620 
   621 definition
   622   "enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4, a\<^sub>5]"
   623 
   624 definition
   625   "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"
   626 
   627 definition
   628   "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"
   629 
   630 instance proof
   631 qed (simp_all only: enum_finite_5_def enum_all_finite_5_def enum_ex_finite_5_def UNIV_finite_5, simp_all)
   632 
   633 end
   634 
   635 hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4 a\<^sub>5
   636 
   637 
   638 subsection {* Closing up *}
   639 
   640 hide_type (open) finite_1 finite_2 finite_3 finite_4 finite_5
   641 hide_const (open) enum enum_all enum_ex all_n_lists ex_n_lists ntrancl
   642 
   643 end
   644