--- a/src/HOL/Enum.thy Sat Oct 20 09:09:37 2012 +0200
+++ b/src/HOL/Enum.thy Sat Oct 20 09:12:16 2012 +0200
@@ -74,61 +74,11 @@
by (simp add: enum_ex)
lemma exists1_code[code]: "(\<exists>!x. P x) \<longleftrightarrow> list_ex1 P enum"
-unfolding list_ex1_iff enum_UNIV by auto
+ by (auto simp add: enum_UNIV list_ex1_iff)
subsection {* Default instances *}
-primrec n_lists :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list list" where
- "n_lists 0 xs = [[]]"
- | "n_lists (Suc n) xs = concat (map (\<lambda>ys. map (\<lambda>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 \<in> set (n_lists n xs) \<Longrightarrow> length ys = n"
- by (induct n arbitrary: ys) auto
-
-lemma set_n_lists: "set (n_lists n xs) = {ys. length ys = n \<and> set ys \<subseteq> set xs}"
-proof (rule set_eqI)
- fix ys :: "'a list"
- show "ys \<in> set (n_lists n xs) \<longleftrightarrow> ys \<in> {ys. length ys = n \<and> set ys \<subseteq> set xs}"
- proof -
- have "ys \<in> set (n_lists n xs) \<Longrightarrow> length ys = n"
- by (induct n arbitrary: ys) auto
- moreover have "\<And>x. ys \<in> set (n_lists n xs) \<Longrightarrow> x \<in> set ys \<Longrightarrow> x \<in> set xs"
- by (induct n arbitrary: ys) auto
- moreover have "set ys \<subseteq> set xs \<Longrightarrow> ys \<in> 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 (\<Union>ys\<in>set (n_lists n xs). (\<lambda>y. y # ys) ` set xs)
- = (\<Sum>ys\<in>set (n_lists n xs). card ((\<lambda>y. y # ys) ` set xs))"
- by (rule card_UN_disjoint) auto
- moreover have "\<And>ys. card ((\<lambda>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 "\<dots> = 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 \<Colon> 'a\<Colon>enum list)"
shows "\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y"
@@ -164,7 +114,7 @@
definition
all_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
where
- "all_n_lists P n = (\<forall>xs \<in> set (n_lists n enum). P xs)"
+ "all_n_lists P n = (\<forall>xs \<in> set (List.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)))"
@@ -174,7 +124,7 @@
definition
ex_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
where
- "ex_n_lists P n = (\<exists>xs \<in> set (n_lists n enum). P xs)"
+ "ex_n_lists P n = (\<exists>xs \<in> set (List.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)))"
@@ -186,7 +136,7 @@
begin
definition
- "enum = map (\<lambda>ys. the o map_of (zip (enum\<Colon>'a list) ys)) (n_lists (length (enum\<Colon>'a\<Colon>enum list)) enum)"
+ "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)"
definition
"enum_all P = all_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))"
@@ -258,7 +208,7 @@
end
lemma enum_fun_code [code]: "enum = (let enum_a = (enum \<Colon> 'a\<Colon>{enum, equal} list)
- in map (\<lambda>ys. the o map_of (zip enum_a ys)) (n_lists (length enum_a) enum))"
+ in map (\<lambda>ys. the o map_of (zip enum_a ys)) (List.n_lists (length enum_a) enum))"
by (simp add: enum_fun_def Let_def)
lemma enum_all_fun_code [code]:
@@ -312,25 +262,11 @@
end
-primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> '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 \<times> 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"
+ "enum = List.product enum enum"
definition
"enum_all P = enum_all (%x. enum_all (%y. P (x, y)))"
@@ -404,7 +340,7 @@
begin
definition
- "enum = map (split Char) (product enum enum)"
+ "enum = map (split Char) (List.product enum enum)"
lemma enum_chars [code]:
"enum = chars"
@@ -461,37 +397,6 @@
qed (auto simp add: enum_UNIV enum_option_def, rule option.exhaust, auto intro: simp add: distinct_map enum_distinct)
end
-primrec sublists :: "'a list \<Rightarrow> 'a list list" where
- "sublists [] = [[]]"
- | "sublists (x#xs) = (let xss = sublists xs in map (Cons x) xss @ xss)"
-
-lemma length_sublists:
- "length (sublists xs) = 2 ^ length xs"
- by (induct xs) (simp_all add: Let_def)
-
-lemma sublists_powset:
- "set ` set (sublists xs) = Pow (set xs)"
-proof -
- have aux: "\<And>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)) = 2 ^ card (set xs)" by (rule card_Pow)
- with assms distinct_card [of xs]
- have "card (Pow (set xs)) = 2 ^ 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 set :: (enum) enum
begin
@@ -745,10 +650,11 @@
hide_const (open) a\<^isub>1 a\<^isub>2 a\<^isub>3 a\<^isub>4 a\<^isub>5
+
subsection {* An executable THE operator on finite types *}
definition
- [code del]: "enum_the P = The P"
+ [code del]: "enum_the = The"
lemma [code]:
"The P = (case filter P enum of [x] => x | _ => enum_the P)"
@@ -782,9 +688,10 @@
code_abort enum_the
code_const enum_the (Eval "(fn p => raise Match)")
+
subsection {* Further operations on finite types *}
-lemma Collect_code[code]:
+lemma Collect_code [code]:
"Collect P = set (filter P enum)"
by (auto simp add: enum_UNIV)
@@ -796,11 +703,11 @@
"tranclp r a b \<equiv> (a, b) \<in> trancl {(x, y). r x y}"
by (simp add: trancl_def)
-lemma rtranclp_rtrancl_eq[code, no_atp]:
+lemma rtranclp_rtrancl_eq [code, no_atp]:
"rtranclp r x y = ((x, y) : rtrancl {(x, y). r x y})"
unfolding rtrancl_def by auto
-lemma max_ext_eq[code]:
+lemma max_ext_eq [code]:
"max_ext R = {(X, Y). finite X & finite Y & Y ~={} & (ALL x. x : X --> (EX xa : Y. (x, xa) : R))}"
by (auto simp add: max_ext.simps)
@@ -813,157 +720,26 @@
unfolding mlex_prod_def by auto
subsection {* Executable accessible part *}
-(* FIXME: should be moved somewhere else !? *)
-
-subsubsection {* Finite monotone eventually stable sequences *}
-
-lemma finite_mono_remains_stable_implies_strict_prefix:
- fixes f :: "nat \<Rightarrow> 'a::order"
- assumes S: "finite (range f)" "mono f" and eq: "\<forall>n. f n = f (Suc n) \<longrightarrow> f (Suc n) = f (Suc (Suc n))"
- 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)"
- using assms
-proof -
- have "\<exists>n. f n = f (Suc n)"
- proof (rule ccontr)
- assume "\<not> ?thesis"
- then have "\<And>n. f n \<noteq> f (Suc n)" by auto
- then have "\<And>n. f n < f (Suc n)"
- using `mono f` by (auto simp: le_less mono_iff_le_Suc)
- with lift_Suc_mono_less_iff[of f]
- have "\<And>n m. n < m \<Longrightarrow> f n < f m" by auto
- then have "inj f"
- by (auto simp: inj_on_def) (metis linorder_less_linear order_less_imp_not_eq)
- with `finite (range f)` have "finite (UNIV::nat set)"
- by (rule finite_imageD)
- then show False by simp
- qed
- then obtain n where n: "f n = f (Suc n)" ..
- def N \<equiv> "LEAST n. f n = f (Suc n)"
- have N: "f N = f (Suc N)"
- unfolding N_def using n by (rule LeastI)
- show ?thesis
- proof (intro exI[of _ N] conjI allI impI)
- fix n assume "N \<le> n"
- then have "\<And>m. N \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> f m = f N"
- proof (induct rule: dec_induct)
- case (step n) then show ?case
- using eq[rule_format, of "n - 1"] N
- by (cases n) (auto simp add: le_Suc_eq)
- qed simp
- from this[of n] `N \<le> n` show "f N = f n" by auto
- next
- fix n m :: nat assume "m < n" "n \<le> N"
- then show "f m < f n"
- proof (induct rule: less_Suc_induct[consumes 1])
- case (1 i)
- then have "i < N" by simp
- then have "f i \<noteq> f (Suc i)"
- unfolding N_def by (rule not_less_Least)
- with `mono f` show ?case by (simp add: mono_iff_le_Suc less_le)
- qed auto
- qed
-qed
-
-lemma finite_mono_strict_prefix_implies_finite_fixpoint:
- fixes f :: "nat \<Rightarrow> 'a set"
- assumes S: "\<And>i. f i \<subseteq> S" "finite S"
- 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)"
- shows "f (card S) = (\<Union>n. f n)"
-proof -
- 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
-
- { fix i have "i \<le> N \<Longrightarrow> i \<le> card (f i)"
- proof (induct i)
- case 0 then show ?case by simp
- next
- case (Suc i)
- with inj[rule_format, of "Suc i" i]
- have "(f i) \<subset> (f (Suc i))" by auto
- moreover have "finite (f (Suc i))" using S by (rule finite_subset)
- ultimately have "card (f i) < card (f (Suc i))" by (intro psubset_card_mono)
- with Suc show ?case using inj by auto
- qed
- }
- then have "N \<le> card (f N)" by simp
- also have "\<dots> \<le> card S" using S by (intro card_mono)
- finally have "f (card S) = f N" using eq by auto
- then show ?thesis using eq inj[rule_format, of N]
- apply auto
- apply (case_tac "n < N")
- apply (auto simp: not_less)
- done
-qed
-
-subsubsection {* Bounded accessible part *}
-
-fun bacc :: "('a * 'a) set => nat => 'a set"
-where
- "bacc r 0 = {x. \<forall> y. (y, x) \<notin> r}"
-| "bacc r (Suc n) = (bacc r n Un {x. \<forall> y. (y, x) : r --> y : bacc r n})"
-
-lemma bacc_subseteq_acc:
- "bacc r n \<subseteq> acc r"
-by (induct n) (auto intro: acc.intros)
-
-lemma bacc_mono:
- "n <= m ==> bacc r n \<subseteq> bacc r m"
-by (induct rule: dec_induct) auto
-
-lemma bacc_upper_bound:
- "bacc (r :: ('a * 'a) set) (card (UNIV :: ('a :: enum) set)) = (UN n. bacc r n)"
-proof -
- have "mono (bacc r)" unfolding mono_def by (simp add: bacc_mono)
- moreover have "\<forall>n. bacc r n = bacc r (Suc n) \<longrightarrow> bacc r (Suc n) = bacc r (Suc (Suc n))" by auto
- moreover have "finite (range (bacc r))" by auto
- ultimately show ?thesis
- by (intro finite_mono_strict_prefix_implies_finite_fixpoint)
- (auto intro: finite_mono_remains_stable_implies_strict_prefix simp add: enum_UNIV)
-qed
-
-lemma acc_subseteq_bacc:
- assumes "finite r"
- shows "acc r \<subseteq> (UN n. bacc r n)"
-proof
- fix x
- assume "x : acc r"
- then have "\<exists> n. x : bacc r n"
- proof (induct x arbitrary: rule: acc.induct)
- case (accI x)
- then have "\<forall>y. \<exists> n. (y, x) \<in> r --> y : bacc r n" by simp
- from choice[OF this] obtain n where n: "\<forall>y. (y, x) \<in> r \<longrightarrow> y \<in> bacc r (n y)" ..
- obtain n where "\<And>y. (y, x) : r \<Longrightarrow> y : bacc r n"
- proof
- fix y assume y: "(y, x) : r"
- with n have "y : bacc r (n y)" by auto
- moreover have "n y <= Max ((%(y, x). n y) ` r)"
- using y `finite r` by (auto intro!: Max_ge)
- note bacc_mono[OF this, of r]
- ultimately show "y : bacc r (Max ((%(y, x). n y) ` r))" by auto
- qed
- then show ?case
- by (auto simp add: Let_def intro!: exI[of _ "Suc n"])
- qed
- then show "x : (UN n. bacc r n)" by auto
-qed
-
-lemma acc_bacc_eq: "acc ((set xs) :: (('a :: enum) * 'a) set) = bacc (set xs) (card (UNIV :: 'a set))"
-by (metis acc_subseteq_bacc bacc_subseteq_acc bacc_upper_bound finite_set order_eq_iff)
definition
[code del]: "card_UNIV = card UNIV"
lemma [code]:
"card_UNIV TYPE('a :: enum) = card (set (Enum.enum :: 'a list))"
-unfolding card_UNIV_def enum_UNIV ..
+ unfolding card_UNIV_def enum_UNIV ..
-declare acc_bacc_eq[folded card_UNIV_def, code]
+lemma [code]:
+ fixes xs :: "('a::finite \<times> 'a) list"
+ shows "acc (set xs) = bacc (set xs) (card_UNIV TYPE('a))"
+ by (simp add: card_UNIV_def acc_bacc_eq)
-lemma [code_unfold]: "accp r = (%x. x : acc {(x, y). r x y})"
-unfolding acc_def by simp
+lemma [code_unfold]: "accp r = (\<lambda>x. x \<in> acc {(x, y). r x y})"
+ unfolding acc_def by simp
subsection {* Closing up *}
hide_type (open) finite_1 finite_2 finite_3 finite_4 finite_5
-hide_const (open) enum enum_all enum_ex n_lists all_n_lists ex_n_lists product ntrancl
+hide_const (open) enum enum_all enum_ex all_n_lists ex_n_lists ntrancl
end
+