src/HOL/Enum.thy
changeset 49948 744934b818c7
parent 48123 104e5fccea12
child 49949 be3dd2e602e8
equal deleted inserted replaced
49947:29cd291bfea6 49948:744934b818c7
    72 
    72 
    73 lemma exists_code [code]: "(\<exists>x. P x) \<longleftrightarrow> enum_ex P"
    73 lemma exists_code [code]: "(\<exists>x. P x) \<longleftrightarrow> enum_ex P"
    74   by (simp add: enum_ex)
    74   by (simp add: enum_ex)
    75 
    75 
    76 lemma exists1_code[code]: "(\<exists>!x. P x) \<longleftrightarrow> list_ex1 P enum"
    76 lemma exists1_code[code]: "(\<exists>!x. P x) \<longleftrightarrow> list_ex1 P enum"
    77 unfolding list_ex1_iff enum_UNIV by auto
    77   by (auto simp add: enum_UNIV list_ex1_iff)
    78 
    78 
    79 
    79 
    80 subsection {* Default instances *}
    80 subsection {* Default instances *}
    81 
       
    82 primrec n_lists :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list list" where
       
    83   "n_lists 0 xs = [[]]"
       
    84   | "n_lists (Suc n) xs = concat (map (\<lambda>ys. map (\<lambda>y. y # ys) xs) (n_lists n xs))"
       
    85 
       
    86 lemma n_lists_Nil [simp]: "n_lists n [] = (if n = 0 then [[]] else [])"
       
    87   by (induct n) simp_all
       
    88 
       
    89 lemma length_n_lists: "length (n_lists n xs) = length xs ^ n"
       
    90   by (induct n) (auto simp add: length_concat o_def listsum_triv)
       
    91 
       
    92 lemma length_n_lists_elem: "ys \<in> set (n_lists n xs) \<Longrightarrow> length ys = n"
       
    93   by (induct n arbitrary: ys) auto
       
    94 
       
    95 lemma set_n_lists: "set (n_lists n xs) = {ys. length ys = n \<and> set ys \<subseteq> set xs}"
       
    96 proof (rule set_eqI)
       
    97   fix ys :: "'a list"
       
    98   show "ys \<in> set (n_lists n xs) \<longleftrightarrow> ys \<in> {ys. length ys = n \<and> set ys \<subseteq> set xs}"
       
    99   proof -
       
   100     have "ys \<in> set (n_lists n xs) \<Longrightarrow> length ys = n"
       
   101       by (induct n arbitrary: ys) auto
       
   102     moreover have "\<And>x. ys \<in> set (n_lists n xs) \<Longrightarrow> x \<in> set ys \<Longrightarrow> x \<in> set xs"
       
   103       by (induct n arbitrary: ys) auto
       
   104     moreover have "set ys \<subseteq> set xs \<Longrightarrow> ys \<in> set (n_lists (length ys) xs)"
       
   105       by (induct ys) auto
       
   106     ultimately show ?thesis by auto
       
   107   qed
       
   108 qed
       
   109 
       
   110 lemma distinct_n_lists:
       
   111   assumes "distinct xs"
       
   112   shows "distinct (n_lists n xs)"
       
   113 proof (rule card_distinct)
       
   114   from assms have card_length: "card (set xs) = length xs" by (rule distinct_card)
       
   115   have "card (set (n_lists n xs)) = card (set xs) ^ n"
       
   116   proof (induct n)
       
   117     case 0 then show ?case by simp
       
   118   next
       
   119     case (Suc n)
       
   120     moreover have "card (\<Union>ys\<in>set (n_lists n xs). (\<lambda>y. y # ys) ` set xs)
       
   121       = (\<Sum>ys\<in>set (n_lists n xs). card ((\<lambda>y. y # ys) ` set xs))"
       
   122       by (rule card_UN_disjoint) auto
       
   123     moreover have "\<And>ys. card ((\<lambda>y. y # ys) ` set xs) = card (set xs)"
       
   124       by (rule card_image) (simp add: inj_on_def)
       
   125     ultimately show ?case by auto
       
   126   qed
       
   127   also have "\<dots> = length xs ^ n" by (simp add: card_length)
       
   128   finally show "card (set (n_lists n xs)) = length (n_lists n xs)"
       
   129     by (simp add: length_n_lists)
       
   130 qed
       
   131 
    81 
   132 lemma map_of_zip_enum_is_Some:
    82 lemma map_of_zip_enum_is_Some:
   133   assumes "length ys = length (enum \<Colon> 'a\<Colon>enum list)"
    83   assumes "length ys = length (enum \<Colon> 'a\<Colon>enum list)"
   134   shows "\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y"
    84   shows "\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y"
   135 proof -
    85 proof -
   162 qed
   112 qed
   163 
   113 
   164 definition
   114 definition
   165   all_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
   115   all_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
   166 where
   116 where
   167   "all_n_lists P n = (\<forall>xs \<in> set (n_lists n enum). P xs)"
   117   "all_n_lists P n = (\<forall>xs \<in> set (List.n_lists n enum). P xs)"
   168 
   118 
   169 lemma [code]:
   119 lemma [code]:
   170   "all_n_lists P n = (if n = 0 then P [] else enum_all (%x. all_n_lists (%xs. P (x # xs)) (n - 1)))"
   120   "all_n_lists P n = (if n = 0 then P [] else enum_all (%x. all_n_lists (%xs. P (x # xs)) (n - 1)))"
   171 unfolding all_n_lists_def enum_all
   121 unfolding all_n_lists_def enum_all
   172 by (cases n) (auto simp add: enum_UNIV)
   122 by (cases n) (auto simp add: enum_UNIV)
   173 
   123 
   174 definition
   124 definition
   175   ex_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
   125   ex_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
   176 where
   126 where
   177   "ex_n_lists P n = (\<exists>xs \<in> set (n_lists n enum). P xs)"
   127   "ex_n_lists P n = (\<exists>xs \<in> set (List.n_lists n enum). P xs)"
   178 
   128 
   179 lemma [code]:
   129 lemma [code]:
   180   "ex_n_lists P n = (if n = 0 then P [] else enum_ex (%x. ex_n_lists (%xs. P (x # xs)) (n - 1)))"
   130   "ex_n_lists P n = (if n = 0 then P [] else enum_ex (%x. ex_n_lists (%xs. P (x # xs)) (n - 1)))"
   181 unfolding ex_n_lists_def enum_ex
   131 unfolding ex_n_lists_def enum_ex
   182 by (cases n) (auto simp add: enum_UNIV)
   132 by (cases n) (auto simp add: enum_UNIV)
   184 
   134 
   185 instantiation "fun" :: (enum, enum) enum
   135 instantiation "fun" :: (enum, enum) enum
   186 begin
   136 begin
   187 
   137 
   188 definition
   138 definition
   189   "enum = map (\<lambda>ys. the o map_of (zip (enum\<Colon>'a list) ys)) (n_lists (length (enum\<Colon>'a\<Colon>enum list)) enum)"
   139   "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)"
   190 
   140 
   191 definition
   141 definition
   192   "enum_all P = all_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))"
   142   "enum_all P = all_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))"
   193 
   143 
   194 definition
   144 definition
   256 qed
   206 qed
   257 
   207 
   258 end
   208 end
   259 
   209 
   260 lemma enum_fun_code [code]: "enum = (let enum_a = (enum \<Colon> 'a\<Colon>{enum, equal} list)
   210 lemma enum_fun_code [code]: "enum = (let enum_a = (enum \<Colon> 'a\<Colon>{enum, equal} list)
   261   in map (\<lambda>ys. the o map_of (zip enum_a ys)) (n_lists (length enum_a) enum))"
   211   in map (\<lambda>ys. the o map_of (zip enum_a ys)) (List.n_lists (length enum_a) enum))"
   262   by (simp add: enum_fun_def Let_def)
   212   by (simp add: enum_fun_def Let_def)
   263 
   213 
   264 lemma enum_all_fun_code [code]:
   214 lemma enum_all_fun_code [code]:
   265   "enum_all P = (let enum_a = (enum :: 'a::{enum, equal} list)
   215   "enum_all P = (let enum_a = (enum :: 'a::{enum, equal} list)
   266    in all_n_lists (\<lambda>bs. P (the o map_of (zip enum_a bs))) (length enum_a))"
   216    in all_n_lists (\<lambda>bs. P (the o map_of (zip enum_a bs))) (length enum_a))"
   310     unfolding enum_ex_bool_def by (auto, case_tac x) auto
   260     unfolding enum_ex_bool_def by (auto, case_tac x) auto
   311 qed (auto simp add: enum_bool_def UNIV_bool)
   261 qed (auto simp add: enum_bool_def UNIV_bool)
   312 
   262 
   313 end
   263 end
   314 
   264 
   315 primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
       
   316   "product [] _ = []"
       
   317   | "product (x#xs) ys = map (Pair x) ys @ product xs ys"
       
   318 
       
   319 lemma product_list_set:
       
   320   "set (product xs ys) = set xs \<times> set ys"
       
   321   by (induct xs) auto
       
   322 
       
   323 lemma distinct_product:
       
   324   assumes "distinct xs" and "distinct ys"
       
   325   shows "distinct (product xs ys)"
       
   326   using assms by (induct xs)
       
   327     (auto intro: inj_onI simp add: product_list_set distinct_map)
       
   328 
       
   329 instantiation prod :: (enum, enum) enum
   265 instantiation prod :: (enum, enum) enum
   330 begin
   266 begin
   331 
   267 
   332 definition
   268 definition
   333   "enum = product enum enum"
   269   "enum = List.product enum enum"
   334 
   270 
   335 definition
   271 definition
   336   "enum_all P = enum_all (%x. enum_all (%y. P (x, y)))"
   272   "enum_all P = enum_all (%x. enum_all (%y. P (x, y)))"
   337 
   273 
   338 definition
   274 definition
   402 
   338 
   403 instantiation char :: enum
   339 instantiation char :: enum
   404 begin
   340 begin
   405 
   341 
   406 definition
   342 definition
   407   "enum = map (split Char) (product enum enum)"
   343   "enum = map (split Char) (List.product enum enum)"
   408 
   344 
   409 lemma enum_chars [code]:
   345 lemma enum_chars [code]:
   410   "enum = chars"
   346   "enum = chars"
   411   unfolding enum_char_def chars_def enum_nibble_def by simp
   347   unfolding enum_char_def chars_def enum_nibble_def by simp
   412 
   348 
   458   show "enum_ex (P :: 'a option \<Rightarrow> bool) = (\<exists>x. P x)"
   394   show "enum_ex (P :: 'a option \<Rightarrow> bool) = (\<exists>x. P x)"
   459     unfolding enum_ex_option_def enum_ex
   395     unfolding enum_ex_option_def enum_ex
   460     by (auto, case_tac x) auto
   396     by (auto, case_tac x) auto
   461 qed (auto simp add: enum_UNIV enum_option_def, rule option.exhaust, auto intro: simp add: distinct_map enum_distinct)
   397 qed (auto simp add: enum_UNIV enum_option_def, rule option.exhaust, auto intro: simp add: distinct_map enum_distinct)
   462 end
   398 end
   463 
       
   464 primrec sublists :: "'a list \<Rightarrow> 'a list list" where
       
   465   "sublists [] = [[]]"
       
   466   | "sublists (x#xs) = (let xss = sublists xs in map (Cons x) xss @ xss)"
       
   467 
       
   468 lemma length_sublists:
       
   469   "length (sublists xs) = 2 ^ length xs"
       
   470   by (induct xs) (simp_all add: Let_def)
       
   471 
       
   472 lemma sublists_powset:
       
   473   "set ` set (sublists xs) = Pow (set xs)"
       
   474 proof -
       
   475   have aux: "\<And>x A. set ` Cons x ` A = insert x ` set ` A"
       
   476     by (auto simp add: image_def)
       
   477   have "set (map set (sublists xs)) = Pow (set xs)"
       
   478     by (induct xs)
       
   479       (simp_all add: aux Let_def Pow_insert Un_commute comp_def del: map_map)
       
   480   then show ?thesis by simp
       
   481 qed
       
   482 
       
   483 lemma distinct_set_sublists:
       
   484   assumes "distinct xs"
       
   485   shows "distinct (map set (sublists xs))"
       
   486 proof (rule card_distinct)
       
   487   have "finite (set xs)" by rule
       
   488   then have "card (Pow (set xs)) = 2 ^ card (set xs)" by (rule card_Pow)
       
   489   with assms distinct_card [of xs]
       
   490     have "card (Pow (set xs)) = 2 ^ length xs" by simp
       
   491   then show "card (set (map set (sublists xs))) = length (map set (sublists xs))"
       
   492     by (simp add: sublists_powset length_sublists)
       
   493 qed
       
   494 
   399 
   495 instantiation set :: (enum) enum
   400 instantiation set :: (enum) enum
   496 begin
   401 begin
   497 
   402 
   498 definition
   403 definition
   743 
   648 
   744 end
   649 end
   745 
   650 
   746 hide_const (open) a\<^isub>1 a\<^isub>2 a\<^isub>3 a\<^isub>4 a\<^isub>5
   651 hide_const (open) a\<^isub>1 a\<^isub>2 a\<^isub>3 a\<^isub>4 a\<^isub>5
   747 
   652 
       
   653 
   748 subsection {* An executable THE operator on finite types *}
   654 subsection {* An executable THE operator on finite types *}
   749 
   655 
   750 definition
   656 definition
   751   [code del]: "enum_the P = The P"
   657   [code del]: "enum_the = The"
   752 
   658 
   753 lemma [code]:
   659 lemma [code]:
   754   "The P = (case filter P enum of [x] => x | _ => enum_the P)"
   660   "The P = (case filter P enum of [x] => x | _ => enum_the P)"
   755 proof -
   661 proof -
   756   {
   662   {
   780 qed
   686 qed
   781 
   687 
   782 code_abort enum_the
   688 code_abort enum_the
   783 code_const enum_the (Eval "(fn p => raise Match)")
   689 code_const enum_the (Eval "(fn p => raise Match)")
   784 
   690 
       
   691 
   785 subsection {* Further operations on finite types *}
   692 subsection {* Further operations on finite types *}
   786 
   693 
   787 lemma Collect_code[code]:
   694 lemma Collect_code [code]:
   788   "Collect P = set (filter P enum)"
   695   "Collect P = set (filter P enum)"
   789 by (auto simp add: enum_UNIV)
   696 by (auto simp add: enum_UNIV)
   790 
   697 
   791 lemma [code]:
   698 lemma [code]:
   792   "Id = image (%x. (x, x)) (set Enum.enum)"
   699   "Id = image (%x. (x, x)) (set Enum.enum)"
   794 
   701 
   795 lemma tranclp_unfold [code, no_atp]:
   702 lemma tranclp_unfold [code, no_atp]:
   796   "tranclp r a b \<equiv> (a, b) \<in> trancl {(x, y). r x y}"
   703   "tranclp r a b \<equiv> (a, b) \<in> trancl {(x, y). r x y}"
   797 by (simp add: trancl_def)
   704 by (simp add: trancl_def)
   798 
   705 
   799 lemma rtranclp_rtrancl_eq[code, no_atp]:
   706 lemma rtranclp_rtrancl_eq [code, no_atp]:
   800   "rtranclp r x y = ((x, y) : rtrancl {(x, y). r x y})"
   707   "rtranclp r x y = ((x, y) : rtrancl {(x, y). r x y})"
   801 unfolding rtrancl_def by auto
   708 unfolding rtrancl_def by auto
   802 
   709 
   803 lemma max_ext_eq[code]:
   710 lemma max_ext_eq [code]:
   804   "max_ext R = {(X, Y). finite X & finite Y & Y ~={} & (ALL x. x : X --> (EX xa : Y. (x, xa) : R))}"
   711   "max_ext R = {(X, Y). finite X & finite Y & Y ~={} & (ALL x. x : X --> (EX xa : Y. (x, xa) : R))}"
   805 by (auto simp add: max_ext.simps)
   712 by (auto simp add: max_ext.simps)
   806 
   713 
   807 lemma max_extp_eq[code]:
   714 lemma max_extp_eq[code]:
   808   "max_extp r x y = ((x, y) : max_ext {(x, y). r x y})"
   715   "max_extp r x y = ((x, y) : max_ext {(x, y). r x y})"
   811 lemma mlex_eq[code]:
   718 lemma mlex_eq[code]:
   812   "f <*mlex*> R = {(x, y). f x < f y \<or> (f x <= f y \<and> (x, y) : R)}"
   719   "f <*mlex*> R = {(x, y). f x < f y \<or> (f x <= f y \<and> (x, y) : R)}"
   813 unfolding mlex_prod_def by auto
   720 unfolding mlex_prod_def by auto
   814 
   721 
   815 subsection {* Executable accessible part *}
   722 subsection {* Executable accessible part *}
   816 (* FIXME: should be moved somewhere else !? *)
       
   817 
       
   818 subsubsection {* Finite monotone eventually stable sequences *}
       
   819 
       
   820 lemma finite_mono_remains_stable_implies_strict_prefix:
       
   821   fixes f :: "nat \<Rightarrow> 'a::order"
       
   822   assumes S: "finite (range f)" "mono f" and eq: "\<forall>n. f n = f (Suc n) \<longrightarrow> f (Suc n) = f (Suc (Suc n))"
       
   823   shows "\<exists>N. (\<forall>n\<le>N. \<forall>m\<le>N. m < n \<longrightarrow> f m < f n) \<and> (\<forall>n\<ge>N. f N = f n)"
       
   824   using assms
       
   825 proof -
       
   826   have "\<exists>n. f n = f (Suc n)"
       
   827   proof (rule ccontr)
       
   828     assume "\<not> ?thesis"
       
   829     then have "\<And>n. f n \<noteq> f (Suc n)" by auto
       
   830     then have "\<And>n. f n < f (Suc n)"
       
   831       using  `mono f` by (auto simp: le_less mono_iff_le_Suc)
       
   832     with lift_Suc_mono_less_iff[of f]
       
   833     have "\<And>n m. n < m \<Longrightarrow> f n < f m" by auto
       
   834     then have "inj f"
       
   835       by (auto simp: inj_on_def) (metis linorder_less_linear order_less_imp_not_eq)
       
   836     with `finite (range f)` have "finite (UNIV::nat set)"
       
   837       by (rule finite_imageD)
       
   838     then show False by simp
       
   839   qed
       
   840   then obtain n where n: "f n = f (Suc n)" ..
       
   841   def N \<equiv> "LEAST n. f n = f (Suc n)"
       
   842   have N: "f N = f (Suc N)"
       
   843     unfolding N_def using n by (rule LeastI)
       
   844   show ?thesis
       
   845   proof (intro exI[of _ N] conjI allI impI)
       
   846     fix n assume "N \<le> n"
       
   847     then have "\<And>m. N \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m = f N"
       
   848     proof (induct rule: dec_induct)
       
   849       case (step n) then show ?case
       
   850         using eq[rule_format, of "n - 1"] N
       
   851         by (cases n) (auto simp add: le_Suc_eq)
       
   852     qed simp
       
   853     from this[of n] `N \<le> n` show "f N = f n" by auto
       
   854   next
       
   855     fix n m :: nat assume "m < n" "n \<le> N"
       
   856     then show "f m < f n"
       
   857     proof (induct rule: less_Suc_induct[consumes 1])
       
   858       case (1 i)
       
   859       then have "i < N" by simp
       
   860       then have "f i \<noteq> f (Suc i)"
       
   861         unfolding N_def by (rule not_less_Least)
       
   862       with `mono f` show ?case by (simp add: mono_iff_le_Suc less_le)
       
   863     qed auto
       
   864   qed
       
   865 qed
       
   866 
       
   867 lemma finite_mono_strict_prefix_implies_finite_fixpoint:
       
   868   fixes f :: "nat \<Rightarrow> 'a set"
       
   869   assumes S: "\<And>i. f i \<subseteq> S" "finite S"
       
   870     and inj: "\<exists>N. (\<forall>n\<le>N. \<forall>m\<le>N. m < n \<longrightarrow> f m \<subset> f n) \<and> (\<forall>n\<ge>N. f N = f n)"
       
   871   shows "f (card S) = (\<Union>n. f n)"
       
   872 proof -
       
   873   from inj obtain N where inj: "(\<forall>n\<le>N. \<forall>m\<le>N. m < n \<longrightarrow> f m \<subset> f n)" and eq: "(\<forall>n\<ge>N. f N = f n)" by auto
       
   874 
       
   875   { fix i have "i \<le> N \<Longrightarrow> i \<le> card (f i)"
       
   876     proof (induct i)
       
   877       case 0 then show ?case by simp
       
   878     next
       
   879       case (Suc i)
       
   880       with inj[rule_format, of "Suc i" i]
       
   881       have "(f i) \<subset> (f (Suc i))" by auto
       
   882       moreover have "finite (f (Suc i))" using S by (rule finite_subset)
       
   883       ultimately have "card (f i) < card (f (Suc i))" by (intro psubset_card_mono)
       
   884       with Suc show ?case using inj by auto
       
   885     qed
       
   886   }
       
   887   then have "N \<le> card (f N)" by simp
       
   888   also have "\<dots> \<le> card S" using S by (intro card_mono)
       
   889   finally have "f (card S) = f N" using eq by auto
       
   890   then show ?thesis using eq inj[rule_format, of N]
       
   891     apply auto
       
   892     apply (case_tac "n < N")
       
   893     apply (auto simp: not_less)
       
   894     done
       
   895 qed
       
   896 
       
   897 subsubsection {* Bounded accessible part *}
       
   898 
       
   899 fun bacc :: "('a * 'a) set => nat => 'a set" 
       
   900 where
       
   901   "bacc r 0 = {x. \<forall> y. (y, x) \<notin> r}"
       
   902 | "bacc r (Suc n) = (bacc r n Un {x. \<forall> y. (y, x) : r --> y : bacc r n})"
       
   903 
       
   904 lemma bacc_subseteq_acc:
       
   905   "bacc r n \<subseteq> acc r"
       
   906 by (induct n) (auto intro: acc.intros)
       
   907 
       
   908 lemma bacc_mono:
       
   909   "n <= m ==> bacc r n \<subseteq> bacc r m"
       
   910 by (induct rule: dec_induct) auto
       
   911   
       
   912 lemma bacc_upper_bound:
       
   913   "bacc (r :: ('a * 'a) set)  (card (UNIV :: ('a :: enum) set)) = (UN n. bacc r n)"
       
   914 proof -
       
   915   have "mono (bacc r)" unfolding mono_def by (simp add: bacc_mono)
       
   916   moreover have "\<forall>n. bacc r n = bacc r (Suc n) \<longrightarrow> bacc r (Suc n) = bacc r (Suc (Suc n))" by auto
       
   917   moreover have "finite (range (bacc r))" by auto
       
   918   ultimately show ?thesis
       
   919    by (intro finite_mono_strict_prefix_implies_finite_fixpoint)
       
   920      (auto intro: finite_mono_remains_stable_implies_strict_prefix  simp add: enum_UNIV)
       
   921 qed
       
   922 
       
   923 lemma acc_subseteq_bacc:
       
   924   assumes "finite r"
       
   925   shows "acc r \<subseteq> (UN n. bacc r n)"
       
   926 proof
       
   927   fix x
       
   928   assume "x : acc r"
       
   929   then have "\<exists> n. x : bacc r n"
       
   930   proof (induct x arbitrary: rule: acc.induct)
       
   931     case (accI x)
       
   932     then have "\<forall>y. \<exists> n. (y, x) \<in> r --> y : bacc r n" by simp
       
   933     from choice[OF this] obtain n where n: "\<forall>y. (y, x) \<in> r \<longrightarrow> y \<in> bacc r (n y)" ..
       
   934     obtain n where "\<And>y. (y, x) : r \<Longrightarrow> y : bacc r n"
       
   935     proof
       
   936       fix y assume y: "(y, x) : r"
       
   937       with n have "y : bacc r (n y)" by auto
       
   938       moreover have "n y <= Max ((%(y, x). n y) ` r)"
       
   939         using y `finite r` by (auto intro!: Max_ge)
       
   940       note bacc_mono[OF this, of r]
       
   941       ultimately show "y : bacc r (Max ((%(y, x). n y) ` r))" by auto
       
   942     qed
       
   943     then show ?case
       
   944       by (auto simp add: Let_def intro!: exI[of _ "Suc n"])
       
   945   qed
       
   946   then show "x : (UN n. bacc r n)" by auto
       
   947 qed
       
   948 
       
   949 lemma acc_bacc_eq: "acc ((set xs) :: (('a :: enum) * 'a) set) = bacc (set xs) (card (UNIV :: 'a set))"
       
   950 by (metis acc_subseteq_bacc bacc_subseteq_acc bacc_upper_bound finite_set order_eq_iff)
       
   951 
   723 
   952 definition 
   724 definition 
   953   [code del]: "card_UNIV = card UNIV"
   725   [code del]: "card_UNIV = card UNIV"
   954 
   726 
   955 lemma [code]:
   727 lemma [code]:
   956   "card_UNIV TYPE('a :: enum) = card (set (Enum.enum :: 'a list))"
   728   "card_UNIV TYPE('a :: enum) = card (set (Enum.enum :: 'a list))"
   957 unfolding card_UNIV_def enum_UNIV ..
   729   unfolding card_UNIV_def enum_UNIV ..
   958 
   730 
   959 declare acc_bacc_eq[folded card_UNIV_def, code]
   731 lemma [code]:
   960 
   732   fixes xs :: "('a::finite \<times> 'a) list"
   961 lemma [code_unfold]: "accp r = (%x. x : acc {(x, y). r x y})"
   733   shows "acc (set xs) = bacc (set xs) (card_UNIV TYPE('a))"
   962 unfolding acc_def by simp
   734   by (simp add: card_UNIV_def acc_bacc_eq)
       
   735 
       
   736 lemma [code_unfold]: "accp r = (\<lambda>x. x \<in> acc {(x, y). r x y})"
       
   737   unfolding acc_def by simp
   963 
   738 
   964 subsection {* Closing up *}
   739 subsection {* Closing up *}
   965 
   740 
   966 hide_type (open) finite_1 finite_2 finite_3 finite_4 finite_5
   741 hide_type (open) finite_1 finite_2 finite_3 finite_4 finite_5
   967 hide_const (open) enum enum_all enum_ex n_lists all_n_lists ex_n_lists product ntrancl
   742 hide_const (open) enum enum_all enum_ex all_n_lists ex_n_lists ntrancl
   968 
   743 
   969 end
   744 end
       
   745