diff -r fd6f41d349ef -r 051251fde456 src/HOL/Enum.thy --- a/src/HOL/Enum.thy Wed Dec 08 14:25:07 2010 +0100 +++ b/src/HOL/Enum.thy Wed Dec 08 14:25:08 2010 +0100 @@ -10,24 +10,28 @@ class enum = fixes enum :: "'a list" + fixes enum_all :: "('a \ bool) \ bool" + fixes enum_ex :: "('a \ bool) \ bool" assumes UNIV_enum: "UNIV = set enum" and enum_distinct: "distinct enum" + assumes enum_all : "enum_all P = (\ x. P x)" + assumes enum_ex : "enum_ex P = (\ x. P x)" begin subclass finite proof qed (simp add: UNIV_enum) -lemma enum_all: "set enum = UNIV" unfolding UNIV_enum .. +lemma enum_UNIV: "set enum = UNIV" unfolding UNIV_enum .. lemma in_enum: "x \ set enum" - unfolding enum_all by auto + unfolding enum_UNIV 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 + with enum_UNIV show ?thesis by simp qed end @@ -42,13 +46,13 @@ "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) +qed (simp_all add: equal_fun_def enum_UNIV fun_eq_iff) end lemma [code]: - "HOL.equal f g \ list_all (%x. f x = g x) enum" -by (auto simp add: list_all_iff enum_all equal fun_eq_iff) + "HOL.equal f g \ enum_all (%x. f x = g x)" +by (auto simp add: equal enum_all fun_eq_iff) lemma [code nbe]: "HOL.equal (f :: _ \ _) f \ True" @@ -56,21 +60,21 @@ 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) + shows "f \ g \ enum_all (\x. f x \ g x)" + and "f < g \ f \ g \ enum_ex (\x. f x \ g x)" + by (simp_all add: enum_all enum_ex 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 all_code [code]: "(\x. P x) \ enum_all P" + by (simp add: enum_all) -lemma exists_code [code]: "(\x. P x) \ list_ex P enum" - by (simp add: list_ex_iff enum_all) +lemma exists_code [code]: "(\x. P x) \ enum_ex P" + by (simp add: enum_ex) lemma exists1_code[code]: "(\!x. P x) \ list_ex1 P enum" -unfolding list_ex1_iff enum_all by auto +unfolding list_ex1_iff enum_UNIV by auto subsection {* Default instances *} @@ -132,7 +136,7 @@ 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 + then show ?thesis using enum_UNIV by auto qed lemma map_of_zip_enum_inject: @@ -156,12 +160,40 @@ with length enum_distinct show "xs = ys" by (rule map_of_zip_inject) qed +definition + all_n_lists :: "(('a :: enum) list \ bool) \ nat \ bool" +where + "all_n_lists P n = (\xs \ set (n_lists n enum). P xs)" + +lemma [code]: + "all_n_lists P n = (if n = 0 then P [] else enum_all (%x. all_n_lists (%xs. P (x # xs)) (n - 1)))" +unfolding all_n_lists_def enum_all +by (cases n) (auto simp add: enum_UNIV) + +definition + ex_n_lists :: "(('a :: enum) list \ bool) \ nat \ bool" +where + "ex_n_lists P n = (\xs \ set (n_lists n enum). P xs)" + +lemma [code]: + "ex_n_lists P n = (if n = 0 then P [] else enum_ex (%x. ex_n_lists (%xs. P (x # xs)) (n - 1)))" +unfolding ex_n_lists_def enum_ex +by (cases n) (auto simp add: enum_UNIV) + + 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)" +definition + "enum_all P = all_n_lists (\bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))" + +definition + "enum_ex P = ex_n_lists (\bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))" + + instance proof show "UNIV = set (enum \ ('a \ 'b) list)" proof (rule UNIV_eq_I) @@ -176,6 +208,50 @@ 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) +next + fix P + show "enum_all (P :: ('a \ 'b) \ bool) = (\x. P x)" + proof + assume "enum_all P" + show "\x. P x" + proof + fix f :: "'a \ 'b" + have f: "f = the \ map_of (zip (enum \ 'a\enum list) (map f enum))" + by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) + from `enum_all P` have "P (the \ map_of (zip enum (map f enum)))" + unfolding enum_all_fun_def all_n_lists_def + apply (simp add: set_n_lists) + apply (erule_tac x="map f enum" in allE) + apply (auto intro!: in_enum) + done + from this f show "P f" by auto + qed + next + assume "\x. P x" + from this show "enum_all P" + unfolding enum_all_fun_def all_n_lists_def by auto + qed +next + fix P + show "enum_ex (P :: ('a \ 'b) \ bool) = (\x. P x)" + proof + assume "enum_ex P" + from this show "\x. P x" + unfolding enum_ex_fun_def ex_n_lists_def by auto + next + assume "\x. P x" + from this obtain f where "P f" .. + have f: "f = the \ map_of (zip (enum \ 'a\enum list) (map f enum))" + by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) + from `P f` this have "P (the \ map_of (zip (enum \ 'a\enum list) (map f enum)))" + by auto + from this show "enum_ex P" + unfolding enum_ex_fun_def ex_n_lists_def + apply (auto simp add: set_n_lists) + apply (rule_tac x="map f enum" in exI) + apply (auto intro!: in_enum) + done + qed qed end @@ -184,14 +260,30 @@ 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) +lemma enum_all_fun_code [code]: + "enum_all P = (let enum_a = (enum :: 'a::{enum, equal} list) + in all_n_lists (\bs. P (the o map_of (zip enum_a bs))) (length enum_a))" + by (simp add: enum_all_fun_def Let_def) + +lemma enum_ex_fun_code [code]: + "enum_ex P = (let enum_a = (enum :: 'a::{enum, equal} list) + in ex_n_lists (\bs. P (the o map_of (zip enum_a bs))) (length enum_a))" + by (simp add: enum_ex_fun_def Let_def) + instantiation unit :: enum begin definition "enum = [()]" +definition + "enum_all P = P ()" + +definition + "enum_ex P = P ()" + instance proof -qed (simp_all add: enum_unit_def UNIV_unit) +qed (auto simp add: enum_unit_def UNIV_unit enum_all_unit_def enum_ex_unit_def intro: unit.exhaust) end @@ -201,8 +293,21 @@ definition "enum = [False, True]" +definition + "enum_all P = (P False \ P True)" + +definition + "enum_ex P = (P False \ P True)" + instance proof -qed (simp_all add: enum_bool_def UNIV_bool) + fix P + show "enum_all (P :: bool \ bool) = (\x. P x)" + unfolding enum_all_bool_def by (auto, case_tac x) auto +next + fix P + show "enum_ex (P :: bool \ bool) = (\x. P x)" + unfolding enum_ex_bool_def by (auto, case_tac x) auto +qed (auto simp add: enum_bool_def UNIV_bool) end @@ -226,8 +331,16 @@ definition "enum = product enum enum" +definition + "enum_all P = enum_all (%x. enum_all (%y. P (x, y)))" + +definition + "enum_ex P = enum_ex (%x. enum_ex (%y. P (x, y)))" + + instance by default - (simp_all add: enum_prod_def product_list_set distinct_product enum_all enum_distinct) + (simp_all add: enum_prod_def product_list_set distinct_product + enum_UNIV enum_distinct enum_all_prod_def enum_all enum_ex_prod_def enum_ex) end @@ -237,8 +350,23 @@ 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) +definition + "enum_all P = (enum_all (%x. P (Inl x)) \ enum_all (%x. P (Inr x)))" + +definition + "enum_ex P = (enum_ex (%x. P (Inl x)) \ enum_ex (%x. P (Inr x)))" + +instance proof + fix P + show "enum_all (P :: ('a + 'b) \ bool) = (\x. P x)" + unfolding enum_all_sum_def enum_all + by (auto, case_tac x) auto +next + fix P + show "enum_ex (P :: ('a + 'b) \ bool) = (\x. P x)" + unfolding enum_ex_sum_def enum_ex + by (auto, case_tac x) auto +qed (auto simp add: enum_UNIV enum_sum_def, case_tac x, auto intro: inj_onI simp add: distinct_map enum_distinct) end @@ -280,7 +408,24 @@ "enum = [Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7, Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF]" +definition + "enum_all P = (P Nibble0 \ P Nibble1 \ P Nibble2 \ P Nibble3 \ P Nibble4 \ P Nibble5 \ P Nibble6 \ P Nibble7 + \ P Nibble8 \ P Nibble9 \ P NibbleA \ P NibbleB \ P NibbleC \ P NibbleD \ P NibbleE \ P NibbleF)" + +definition + "enum_ex P = (P Nibble0 \ P Nibble1 \ P Nibble2 \ P Nibble3 \ P Nibble4 \ P Nibble5 \ P Nibble6 \ P Nibble7 + \ P Nibble8 \ P Nibble9 \ P NibbleA \ P NibbleB \ P NibbleC \ P NibbleD \ P NibbleE \ P NibbleF)" + instance proof + fix P + show "enum_all (P :: nibble \ bool) = (\x. P x)" + unfolding enum_all_nibble_def + by (auto, case_tac x) auto +next + fix P + show "enum_ex (P :: nibble \ bool) = (\x. P x)" + unfolding enum_ex_nibble_def + by (auto, case_tac x) auto qed (simp_all add: enum_nibble_def UNIV_nibble) end @@ -295,9 +440,29 @@ "enum = chars" unfolding enum_char_def chars_def enum_nibble_def by simp +definition + "enum_all P = list_all P chars" + +definition + "enum_ex P = list_ex P chars" + +lemma set_enum_char: "set (enum :: char list) = UNIV" + by (auto intro: char.exhaust simp add: enum_char_def product_list_set enum_UNIV full_SetCompr_eq [symmetric]) + 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) + fix P + show "enum_all (P :: char \ bool) = (\x. P x)" + unfolding enum_all_char_def enum_chars[symmetric] + by (auto simp add: list_all_iff set_enum_char) +next + fix P + show "enum_ex (P :: char \ bool) = (\x. P x)" + unfolding enum_ex_char_def enum_chars[symmetric] + by (auto simp add: list_ex_iff set_enum_char) +next + show "distinct (enum :: char list)" + by (auto intro: inj_onI simp add: enum_char_def product_list_set distinct_map distinct_product enum_distinct) +qed (auto simp add: set_enum_char) end @@ -307,8 +472,23 @@ definition "enum = None # map Some enum" +definition + "enum_all P = (P None \ enum_all (%x. P (Some x)))" + +definition + "enum_ex P = (P None \ enum_ex (%x. P (Some x)))" + instance proof -qed (auto simp add: enum_all enum_option_def, rule option.exhaust, auto intro: simp add: distinct_map enum_distinct) + fix P + show "enum_all (P :: 'a option \ bool) = (\x. P x)" + unfolding enum_all_option_def enum_all + by (auto, case_tac x) auto +next + fix P + show "enum_ex (P :: 'a option \ bool) = (\x. P x)" + unfolding enum_ex_option_def enum_ex + by (auto, case_tac x) auto +qed (auto simp add: enum_UNIV enum_option_def, rule option.exhaust, auto intro: simp add: distinct_map enum_distinct) end @@ -326,7 +506,22 @@ definition "enum = [a\<^isub>1]" +definition + "enum_all P = P a\<^isub>1" + +definition + "enum_ex P = P a\<^isub>1" + instance proof + fix P + show "enum_all (P :: finite_1 \ bool) = (\x. P x)" + unfolding enum_all_finite_1_def + by (auto, case_tac x) auto +next + fix P + show "enum_ex (P :: finite_1 \ bool) = (\x. P x)" + unfolding enum_ex_finite_1_def + by (auto, case_tac x) auto qed (auto simp add: enum_finite_1_def intro: finite_1.exhaust) end @@ -363,7 +558,22 @@ definition "enum = [a\<^isub>1, a\<^isub>2]" +definition + "enum_all P = (P a\<^isub>1 \ P a\<^isub>2)" + +definition + "enum_ex P = (P a\<^isub>1 \ P a\<^isub>2)" + instance proof + fix P + show "enum_all (P :: finite_2 \ bool) = (\x. P x)" + unfolding enum_all_finite_2_def + by (auto, case_tac x) auto +next + fix P + show "enum_ex (P :: finite_2 \ bool) = (\x. P x)" + unfolding enum_ex_finite_2_def + by (auto, case_tac x) auto qed (auto simp add: enum_finite_2_def intro: finite_2.exhaust) end @@ -403,7 +613,22 @@ definition "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3]" +definition + "enum_all P = (P a\<^isub>1 \ P a\<^isub>2 \ P a\<^isub>3)" + +definition + "enum_ex P = (P a\<^isub>1 \ P a\<^isub>2 \ P a\<^isub>3)" + instance proof + fix P + show "enum_all (P :: finite_3 \ bool) = (\x. P x)" + unfolding enum_all_finite_3_def + by (auto, case_tac x) auto +next + fix P + show "enum_ex (P :: finite_3 \ bool) = (\x. P x)" + unfolding enum_ex_finite_3_def + by (auto, case_tac x) auto qed (auto simp add: enum_finite_3_def intro: finite_3.exhaust) end @@ -442,7 +667,22 @@ definition "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4]" +definition + "enum_all P = (P a\<^isub>1 \ P a\<^isub>2 \ P a\<^isub>3 \ P a\<^isub>4)" + +definition + "enum_ex P = (P a\<^isub>1 \ P a\<^isub>2 \ P a\<^isub>3 \ P a\<^isub>4)" + instance proof + fix P + show "enum_all (P :: finite_4 \ bool) = (\x. P x)" + unfolding enum_all_finite_4_def + by (auto, case_tac x) auto +next + fix P + show "enum_ex (P :: finite_4 \ bool) = (\x. P x)" + unfolding enum_ex_finite_4_def + by (auto, case_tac x) auto qed (auto simp add: enum_finite_4_def intro: finite_4.exhaust) end @@ -464,7 +704,22 @@ definition "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4, a\<^isub>5]" +definition + "enum_all P = (P a\<^isub>1 \ P a\<^isub>2 \ P a\<^isub>3 \ P a\<^isub>4 \ P a\<^isub>5)" + +definition + "enum_ex P = (P a\<^isub>1 \ P a\<^isub>2 \ P a\<^isub>3 \ P a\<^isub>4 \ P a\<^isub>5)" + instance proof + fix P + show "enum_all (P :: finite_5 \ bool) = (\x. P x)" + unfolding enum_all_finite_5_def + by (auto, case_tac x) auto +next + fix P + show "enum_ex (P :: finite_5 \ bool) = (\x. P x)" + unfolding enum_ex_finite_5_def + by (auto, case_tac x) auto qed (auto simp add: enum_finite_5_def intro: finite_5.exhaust) end @@ -473,6 +728,6 @@ hide_type finite_1 finite_2 finite_3 finite_4 finite_5 -hide_const (open) enum n_lists product +hide_const (open) enum enum_all enum_ex n_lists all_n_lists ex_n_lists product end