# HG changeset patch # User bulwahn # Date 1290422095 -3600 # Node ID dc1b5aa908ff23282b42b4ec35d305728f8a4cef # Parent 1598ec648b0d5117e9e9ee34669e934b9504ae46 moving Enum theory from HOL/Library to HOL diff -r 1598ec648b0d -r dc1b5aa908ff src/HOL/Enum.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Enum.thy Mon Nov 22 11:34:55 2010 +0100 @@ -0,0 +1,379 @@ +(* Author: Florian Haftmann, TU Muenchen *) + +header {* Finite types as explicit enumerations *} + +theory Enum +imports Map Main +begin + +subsection {* Class @{text enum} *} + +class enum = + fixes enum :: "'a list" + assumes UNIV_enum: "UNIV = set enum" + and enum_distinct: "distinct enum" +begin + +subclass finite proof +qed (simp add: UNIV_enum) + +lemma enum_all: "set enum = UNIV" unfolding UNIV_enum .. + +lemma in_enum [intro]: "x \ set enum" + unfolding enum_all by auto + +lemma enum_eq_I: + assumes "\x. x \ set xs" + shows "set enum = set xs" +proof - + from assms UNIV_eq_I have "UNIV = set xs" by auto + with enum_all show ?thesis by simp +qed + +end + + +subsection {* Equality and order on functions *} + +instantiation "fun" :: (enum, equal) equal +begin + +definition + "HOL.equal f g \ (\x \ set enum. f x = g x)" + +instance proof +qed (simp_all add: equal_fun_def enum_all fun_eq_iff) + +end + +lemma [code nbe]: + "HOL.equal (f :: _ \ _) f \ True" + by (fact equal_refl) + +lemma order_fun [code]: + fixes f g :: "'a\enum \ 'b\order" + shows "f \ g \ list_all (\x. f x \ g x) enum" + and "f < g \ f \ g \ list_ex (\x. f x \ g x) enum" + by (simp_all add: list_all_iff list_ex_iff enum_all fun_eq_iff le_fun_def order_less_le) + + +subsection {* Quantifiers *} + +lemma all_code [code]: "(\x. P x) \ list_all P enum" + by (simp add: list_all_iff enum_all) + +lemma exists_code [code]: "(\x. P x) \ list_ex P enum" + by (simp add: list_ex_iff enum_all) + + +subsection {* Default instances *} + +primrec n_lists :: "nat \ 'a list \ 'a list list" where + "n_lists 0 xs = [[]]" + | "n_lists (Suc n) xs = concat (map (\ys. map (\y. y # ys) xs) (n_lists n xs))" + +lemma n_lists_Nil [simp]: "n_lists n [] = (if n = 0 then [[]] else [])" + by (induct n) simp_all + +lemma length_n_lists: "length (n_lists n xs) = length xs ^ n" + by (induct n) (auto simp add: length_concat o_def listsum_triv) + +lemma length_n_lists_elem: "ys \ set (n_lists n xs) \ length ys = n" + by (induct n arbitrary: ys) auto + +lemma set_n_lists: "set (n_lists n xs) = {ys. length ys = n \ set ys \ set xs}" +proof (rule set_eqI) + fix ys :: "'a list" + show "ys \ set (n_lists n xs) \ ys \ {ys. length ys = n \ set ys \ set xs}" + proof - + have "ys \ set (n_lists n xs) \ length ys = n" + by (induct n arbitrary: ys) auto + moreover have "\x. ys \ set (n_lists n xs) \ x \ set ys \ x \ set xs" + by (induct n arbitrary: ys) auto + moreover have "set ys \ set xs \ ys \ set (n_lists (length ys) xs)" + by (induct ys) auto + ultimately show ?thesis by auto + qed +qed + +lemma distinct_n_lists: + assumes "distinct xs" + shows "distinct (n_lists n xs)" +proof (rule card_distinct) + from assms have card_length: "card (set xs) = length xs" by (rule distinct_card) + have "card (set (n_lists n xs)) = card (set xs) ^ n" + proof (induct n) + case 0 then show ?case by simp + next + case (Suc n) + moreover have "card (\ys\set (n_lists n xs). (\y. y # ys) ` set xs) + = (\ys\set (n_lists n xs). card ((\y. y # ys) ` set xs))" + by (rule card_UN_disjoint) auto + moreover have "\ys. card ((\y. y # ys) ` set xs) = card (set xs)" + by (rule card_image) (simp add: inj_on_def) + ultimately show ?case by auto + qed + also have "\ = length xs ^ n" by (simp add: card_length) + finally show "card (set (n_lists n xs)) = length (n_lists n xs)" + by (simp add: length_n_lists) +qed + +lemma map_of_zip_enum_is_Some: + assumes "length ys = length (enum \ 'a\enum list)" + shows "\y. map_of (zip (enum \ 'a\enum list) ys) x = Some y" +proof - + from assms have "x \ set (enum \ 'a\enum list) \ + (\y. map_of (zip (enum \ 'a\enum list) ys) x = Some y)" + by (auto intro!: map_of_zip_is_Some) + then show ?thesis using enum_all by auto +qed + +lemma map_of_zip_enum_inject: + fixes xs ys :: "'b\enum list" + assumes length: "length xs = length (enum \ 'a\enum list)" + "length ys = length (enum \ 'a\enum list)" + and map_of: "the \ map_of (zip (enum \ 'a\enum list) xs) = the \ map_of (zip (enum \ 'a\enum list) ys)" + shows "xs = ys" +proof - + have "map_of (zip (enum \ 'a list) xs) = map_of (zip (enum \ 'a list) ys)" + proof + fix x :: 'a + from length map_of_zip_enum_is_Some obtain y1 y2 + where "map_of (zip (enum \ 'a list) xs) x = Some y1" + and "map_of (zip (enum \ 'a list) ys) x = Some y2" by blast + moreover from map_of have "the (map_of (zip (enum \ 'a\enum list) xs) x) = the (map_of (zip (enum \ 'a\enum list) ys) x)" + by (auto dest: fun_cong) + ultimately show "map_of (zip (enum \ 'a\enum list) xs) x = map_of (zip (enum \ 'a\enum list) ys) x" + by simp + qed + with length enum_distinct show "xs = ys" by (rule map_of_zip_inject) +qed + +instantiation "fun" :: (enum, enum) enum +begin + +definition + "enum = map (\ys. the o map_of (zip (enum\'a list) ys)) (n_lists (length (enum\'a\enum list)) enum)" + +instance proof + show "UNIV = set (enum \ ('a \ 'b) list)" + proof (rule UNIV_eq_I) + fix f :: "'a \ 'b" + have "f = the \ map_of (zip (enum \ 'a\enum list) (map f enum))" + by (auto simp add: map_of_zip_map fun_eq_iff) + then show "f \ set enum" + by (auto simp add: enum_fun_def set_n_lists) + qed +next + from map_of_zip_enum_inject + show "distinct (enum \ ('a \ 'b) list)" + by (auto intro!: inj_onI simp add: enum_fun_def + distinct_map distinct_n_lists enum_distinct set_n_lists enum_all) +qed + +end + +lemma enum_fun_code [code]: "enum = (let enum_a = (enum \ 'a\{enum, equal} list) + in map (\ys. the o map_of (zip enum_a ys)) (n_lists (length enum_a) enum))" + by (simp add: enum_fun_def Let_def) + +instantiation unit :: enum +begin + +definition + "enum = [()]" + +instance proof +qed (simp_all add: enum_unit_def UNIV_unit) + +end + +instantiation bool :: enum +begin + +definition + "enum = [False, True]" + +instance proof +qed (simp_all add: enum_bool_def UNIV_bool) + +end + +primrec product :: "'a list \ 'b list \ ('a \ 'b) list" where + "product [] _ = []" + | "product (x#xs) ys = map (Pair x) ys @ product xs ys" + +lemma product_list_set: + "set (product xs ys) = set xs \ set ys" + by (induct xs) auto + +lemma distinct_product: + assumes "distinct xs" and "distinct ys" + shows "distinct (product xs ys)" + using assms by (induct xs) + (auto intro: inj_onI simp add: product_list_set distinct_map) + +instantiation prod :: (enum, enum) enum +begin + +definition + "enum = product enum enum" + +instance by default + (simp_all add: enum_prod_def product_list_set distinct_product enum_all enum_distinct) + +end + +instantiation sum :: (enum, enum) enum +begin + +definition + "enum = map Inl enum @ map Inr enum" + +instance by default + (auto simp add: enum_all enum_sum_def, case_tac x, auto intro: inj_onI simp add: distinct_map enum_distinct) + +end + +primrec sublists :: "'a list \ 'a list list" where + "sublists [] = [[]]" + | "sublists (x#xs) = (let xss = sublists xs in map (Cons x) xss @ xss)" + +lemma length_sublists: + "length (sublists xs) = Suc (Suc (0\nat)) ^ length xs" + by (induct xs) (simp_all add: Let_def) + +lemma sublists_powset: + "set ` set (sublists xs) = Pow (set xs)" +proof - + have aux: "\x A. set ` Cons x ` A = insert x ` set ` A" + by (auto simp add: image_def) + have "set (map set (sublists xs)) = Pow (set xs)" + by (induct xs) + (simp_all add: aux Let_def Pow_insert Un_commute comp_def del: map_map) + then show ?thesis by simp +qed + +lemma distinct_set_sublists: + assumes "distinct xs" + shows "distinct (map set (sublists xs))" +proof (rule card_distinct) + have "finite (set xs)" by rule + then have "card (Pow (set xs)) = Suc (Suc 0) ^ card (set xs)" by (rule card_Pow) + with assms distinct_card [of xs] + have "card (Pow (set xs)) = Suc (Suc 0) ^ length xs" by simp + then show "card (set (map set (sublists xs))) = length (map set (sublists xs))" + by (simp add: sublists_powset length_sublists) +qed + +instantiation nibble :: enum +begin + +definition + "enum = [Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7, + Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF]" + +instance proof +qed (simp_all add: enum_nibble_def UNIV_nibble) + +end + +instantiation char :: enum +begin + +definition + "enum = map (split Char) (product enum enum)" + +lemma enum_chars [code]: + "enum = chars" + unfolding enum_char_def chars_def enum_nibble_def by simp + +instance proof +qed (auto intro: char.exhaust injI simp add: enum_char_def product_list_set enum_all full_SetCompr_eq [symmetric] + distinct_map distinct_product enum_distinct) + +end + +instantiation option :: (enum) enum +begin + +definition + "enum = None # map Some enum" + +instance proof +qed (auto simp add: enum_all enum_option_def, rule option.exhaust, auto intro: simp add: distinct_map enum_distinct) + +end + +subsection {* Small finite types *} + +text {* We define small finite types for the use in Quickcheck *} + +datatype finite_1 = a\<^isub>1 + +instantiation finite_1 :: enum +begin + +definition + "enum = [a\<^isub>1]" + +instance proof +qed (auto simp add: enum_finite_1_def intro: finite_1.exhaust) + +end + +datatype finite_2 = a\<^isub>1 | a\<^isub>2 + +instantiation finite_2 :: enum +begin + +definition + "enum = [a\<^isub>1, a\<^isub>2]" + +instance proof +qed (auto simp add: enum_finite_2_def intro: finite_2.exhaust) + +end + +datatype finite_3 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 + +instantiation finite_3 :: enum +begin + +definition + "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3]" + +instance proof +qed (auto simp add: enum_finite_3_def intro: finite_3.exhaust) + +end + +datatype finite_4 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 | a\<^isub>4 + +instantiation finite_4 :: enum +begin + +definition + "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4]" + +instance proof +qed (auto simp add: enum_finite_4_def intro: finite_4.exhaust) + +end + +datatype finite_5 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 | a\<^isub>4 | a\<^isub>5 + +instantiation finite_5 :: enum +begin + +definition + "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4, a\<^isub>5]" + +instance proof +qed (auto simp add: enum_finite_5_def intro: finite_5.exhaust) + +end + +hide_type finite_1 finite_2 finite_3 finite_4 finite_5 + +end diff -r 1598ec648b0d -r dc1b5aa908ff src/HOL/Library/Enum.thy --- a/src/HOL/Library/Enum.thy Mon Nov 22 11:34:54 2010 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,379 +0,0 @@ -(* Author: Florian Haftmann, TU Muenchen *) - -header {* Finite types as explicit enumerations *} - -theory Enum -imports Map Main -begin - -subsection {* Class @{text enum} *} - -class enum = - fixes enum :: "'a list" - assumes UNIV_enum: "UNIV = set enum" - and enum_distinct: "distinct enum" -begin - -subclass finite proof -qed (simp add: UNIV_enum) - -lemma enum_all: "set enum = UNIV" unfolding UNIV_enum .. - -lemma in_enum [intro]: "x \ set enum" - unfolding enum_all by auto - -lemma enum_eq_I: - assumes "\x. x \ set xs" - shows "set enum = set xs" -proof - - from assms UNIV_eq_I have "UNIV = set xs" by auto - with enum_all show ?thesis by simp -qed - -end - - -subsection {* Equality and order on functions *} - -instantiation "fun" :: (enum, equal) equal -begin - -definition - "HOL.equal f g \ (\x \ set enum. f x = g x)" - -instance proof -qed (simp_all add: equal_fun_def enum_all fun_eq_iff) - -end - -lemma [code nbe]: - "HOL.equal (f :: _ \ _) f \ True" - by (fact equal_refl) - -lemma order_fun [code]: - fixes f g :: "'a\enum \ 'b\order" - shows "f \ g \ list_all (\x. f x \ g x) enum" - and "f < g \ f \ g \ list_ex (\x. f x \ g x) enum" - by (simp_all add: list_all_iff list_ex_iff enum_all fun_eq_iff le_fun_def order_less_le) - - -subsection {* Quantifiers *} - -lemma all_code [code]: "(\x. P x) \ list_all P enum" - by (simp add: list_all_iff enum_all) - -lemma exists_code [code]: "(\x. P x) \ list_ex P enum" - by (simp add: list_ex_iff enum_all) - - -subsection {* Default instances *} - -primrec n_lists :: "nat \ 'a list \ 'a list list" where - "n_lists 0 xs = [[]]" - | "n_lists (Suc n) xs = concat (map (\ys. map (\y. y # ys) xs) (n_lists n xs))" - -lemma n_lists_Nil [simp]: "n_lists n [] = (if n = 0 then [[]] else [])" - by (induct n) simp_all - -lemma length_n_lists: "length (n_lists n xs) = length xs ^ n" - by (induct n) (auto simp add: length_concat o_def listsum_triv) - -lemma length_n_lists_elem: "ys \ set (n_lists n xs) \ length ys = n" - by (induct n arbitrary: ys) auto - -lemma set_n_lists: "set (n_lists n xs) = {ys. length ys = n \ set ys \ set xs}" -proof (rule set_eqI) - fix ys :: "'a list" - show "ys \ set (n_lists n xs) \ ys \ {ys. length ys = n \ set ys \ set xs}" - proof - - have "ys \ set (n_lists n xs) \ length ys = n" - by (induct n arbitrary: ys) auto - moreover have "\x. ys \ set (n_lists n xs) \ x \ set ys \ x \ set xs" - by (induct n arbitrary: ys) auto - moreover have "set ys \ set xs \ ys \ set (n_lists (length ys) xs)" - by (induct ys) auto - ultimately show ?thesis by auto - qed -qed - -lemma distinct_n_lists: - assumes "distinct xs" - shows "distinct (n_lists n xs)" -proof (rule card_distinct) - from assms have card_length: "card (set xs) = length xs" by (rule distinct_card) - have "card (set (n_lists n xs)) = card (set xs) ^ n" - proof (induct n) - case 0 then show ?case by simp - next - case (Suc n) - moreover have "card (\ys\set (n_lists n xs). (\y. y # ys) ` set xs) - = (\ys\set (n_lists n xs). card ((\y. y # ys) ` set xs))" - by (rule card_UN_disjoint) auto - moreover have "\ys. card ((\y. y # ys) ` set xs) = card (set xs)" - by (rule card_image) (simp add: inj_on_def) - ultimately show ?case by auto - qed - also have "\ = length xs ^ n" by (simp add: card_length) - finally show "card (set (n_lists n xs)) = length (n_lists n xs)" - by (simp add: length_n_lists) -qed - -lemma map_of_zip_enum_is_Some: - assumes "length ys = length (enum \ 'a\enum list)" - shows "\y. map_of (zip (enum \ 'a\enum list) ys) x = Some y" -proof - - from assms have "x \ set (enum \ 'a\enum list) \ - (\y. map_of (zip (enum \ 'a\enum list) ys) x = Some y)" - by (auto intro!: map_of_zip_is_Some) - then show ?thesis using enum_all by auto -qed - -lemma map_of_zip_enum_inject: - fixes xs ys :: "'b\enum list" - assumes length: "length xs = length (enum \ 'a\enum list)" - "length ys = length (enum \ 'a\enum list)" - and map_of: "the \ map_of (zip (enum \ 'a\enum list) xs) = the \ map_of (zip (enum \ 'a\enum list) ys)" - shows "xs = ys" -proof - - have "map_of (zip (enum \ 'a list) xs) = map_of (zip (enum \ 'a list) ys)" - proof - fix x :: 'a - from length map_of_zip_enum_is_Some obtain y1 y2 - where "map_of (zip (enum \ 'a list) xs) x = Some y1" - and "map_of (zip (enum \ 'a list) ys) x = Some y2" by blast - moreover from map_of have "the (map_of (zip (enum \ 'a\enum list) xs) x) = the (map_of (zip (enum \ 'a\enum list) ys) x)" - by (auto dest: fun_cong) - ultimately show "map_of (zip (enum \ 'a\enum list) xs) x = map_of (zip (enum \ 'a\enum list) ys) x" - by simp - qed - with length enum_distinct show "xs = ys" by (rule map_of_zip_inject) -qed - -instantiation "fun" :: (enum, enum) enum -begin - -definition - "enum = map (\ys. the o map_of (zip (enum\'a list) ys)) (n_lists (length (enum\'a\enum list)) enum)" - -instance proof - show "UNIV = set (enum \ ('a \ 'b) list)" - proof (rule UNIV_eq_I) - fix f :: "'a \ 'b" - have "f = the \ map_of (zip (enum \ 'a\enum list) (map f enum))" - by (auto simp add: map_of_zip_map fun_eq_iff) - then show "f \ set enum" - by (auto simp add: enum_fun_def set_n_lists) - qed -next - from map_of_zip_enum_inject - show "distinct (enum \ ('a \ 'b) list)" - by (auto intro!: inj_onI simp add: enum_fun_def - distinct_map distinct_n_lists enum_distinct set_n_lists enum_all) -qed - -end - -lemma enum_fun_code [code]: "enum = (let enum_a = (enum \ 'a\{enum, equal} list) - in map (\ys. the o map_of (zip enum_a ys)) (n_lists (length enum_a) enum))" - by (simp add: enum_fun_def Let_def) - -instantiation unit :: enum -begin - -definition - "enum = [()]" - -instance proof -qed (simp_all add: enum_unit_def UNIV_unit) - -end - -instantiation bool :: enum -begin - -definition - "enum = [False, True]" - -instance proof -qed (simp_all add: enum_bool_def UNIV_bool) - -end - -primrec product :: "'a list \ 'b list \ ('a \ 'b) list" where - "product [] _ = []" - | "product (x#xs) ys = map (Pair x) ys @ product xs ys" - -lemma product_list_set: - "set (product xs ys) = set xs \ set ys" - by (induct xs) auto - -lemma distinct_product: - assumes "distinct xs" and "distinct ys" - shows "distinct (product xs ys)" - using assms by (induct xs) - (auto intro: inj_onI simp add: product_list_set distinct_map) - -instantiation prod :: (enum, enum) enum -begin - -definition - "enum = product enum enum" - -instance by default - (simp_all add: enum_prod_def product_list_set distinct_product enum_all enum_distinct) - -end - -instantiation sum :: (enum, enum) enum -begin - -definition - "enum = map Inl enum @ map Inr enum" - -instance by default - (auto simp add: enum_all enum_sum_def, case_tac x, auto intro: inj_onI simp add: distinct_map enum_distinct) - -end - -primrec sublists :: "'a list \ 'a list list" where - "sublists [] = [[]]" - | "sublists (x#xs) = (let xss = sublists xs in map (Cons x) xss @ xss)" - -lemma length_sublists: - "length (sublists xs) = Suc (Suc (0\nat)) ^ length xs" - by (induct xs) (simp_all add: Let_def) - -lemma sublists_powset: - "set ` set (sublists xs) = Pow (set xs)" -proof - - have aux: "\x A. set ` Cons x ` A = insert x ` set ` A" - by (auto simp add: image_def) - have "set (map set (sublists xs)) = Pow (set xs)" - by (induct xs) - (simp_all add: aux Let_def Pow_insert Un_commute comp_def del: map_map) - then show ?thesis by simp -qed - -lemma distinct_set_sublists: - assumes "distinct xs" - shows "distinct (map set (sublists xs))" -proof (rule card_distinct) - have "finite (set xs)" by rule - then have "card (Pow (set xs)) = Suc (Suc 0) ^ card (set xs)" by (rule card_Pow) - with assms distinct_card [of xs] - have "card (Pow (set xs)) = Suc (Suc 0) ^ length xs" by simp - then show "card (set (map set (sublists xs))) = length (map set (sublists xs))" - by (simp add: sublists_powset length_sublists) -qed - -instantiation nibble :: enum -begin - -definition - "enum = [Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7, - Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF]" - -instance proof -qed (simp_all add: enum_nibble_def UNIV_nibble) - -end - -instantiation char :: enum -begin - -definition - "enum = map (split Char) (product enum enum)" - -lemma enum_chars [code]: - "enum = chars" - unfolding enum_char_def chars_def enum_nibble_def by simp - -instance proof -qed (auto intro: char.exhaust injI simp add: enum_char_def product_list_set enum_all full_SetCompr_eq [symmetric] - distinct_map distinct_product enum_distinct) - -end - -instantiation option :: (enum) enum -begin - -definition - "enum = None # map Some enum" - -instance proof -qed (auto simp add: enum_all enum_option_def, rule option.exhaust, auto intro: simp add: distinct_map enum_distinct) - -end - -subsection {* Small finite types *} - -text {* We define small finite types for the use in Quickcheck *} - -datatype finite_1 = a\<^isub>1 - -instantiation finite_1 :: enum -begin - -definition - "enum = [a\<^isub>1]" - -instance proof -qed (auto simp add: enum_finite_1_def intro: finite_1.exhaust) - -end - -datatype finite_2 = a\<^isub>1 | a\<^isub>2 - -instantiation finite_2 :: enum -begin - -definition - "enum = [a\<^isub>1, a\<^isub>2]" - -instance proof -qed (auto simp add: enum_finite_2_def intro: finite_2.exhaust) - -end - -datatype finite_3 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 - -instantiation finite_3 :: enum -begin - -definition - "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3]" - -instance proof -qed (auto simp add: enum_finite_3_def intro: finite_3.exhaust) - -end - -datatype finite_4 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 | a\<^isub>4 - -instantiation finite_4 :: enum -begin - -definition - "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4]" - -instance proof -qed (auto simp add: enum_finite_4_def intro: finite_4.exhaust) - -end - -datatype finite_5 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 | a\<^isub>4 | a\<^isub>5 - -instantiation finite_5 :: enum -begin - -definition - "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4, a\<^isub>5]" - -instance proof -qed (auto simp add: enum_finite_5_def intro: finite_5.exhaust) - -end - -hide_type finite_1 finite_2 finite_3 finite_4 finite_5 - -end