--- a/src/HOL/IsaMakefile Thu Jun 25 16:48:43 2009 +0200
+++ b/src/HOL/IsaMakefile Thu Jun 25 17:07:28 2009 +0200
@@ -319,7 +319,7 @@
Library/Abstract_Rat.thy \
Library/BigO.thy Library/ContNotDenum.thy Library/Efficient_Nat.thy \
Library/Euclidean_Space.thy Library/Sum_Of_Squares.thy Library/positivstellensatz.ML \
- Library/Convex_Euclidean_Space.thy \
+ Library/Code_Set.thy Library/Convex_Euclidean_Space.thy \
Library/sum_of_squares.ML Library/Glbs.thy Library/normarith.ML \
Library/Executable_Set.thy Library/Infinite_Set.thy \
Library/FuncSet.thy Library/Permutations.thy Library/Determinants.thy\
@@ -329,7 +329,7 @@
Library/Fundamental_Theorem_Algebra.thy \
Library/Inner_Product.thy Library/Lattice_Syntax.thy \
Library/Legacy_GCD.thy \
- Library/Library.thy Library/List_Prefix.thy Library/State_Monad.thy \
+ Library/Library.thy Library/List_Prefix.thy Library/List_Set.thy Library/State_Monad.thy \
Library/Nat_Int_Bij.thy Library/Multiset.thy Library/Permutation.thy \
Library/Primes.thy Library/Pocklington.thy Library/Quotient.thy \
Library/Quicksort.thy Library/Nat_Infinity.thy Library/Word.thy \
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Code_Set.thy Thu Jun 25 17:07:28 2009 +0200
@@ -0,0 +1,169 @@
+
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* Executable finite sets *}
+
+theory Code_Set
+imports List_Set
+begin
+
+lemma foldl_apply_inv:
+ assumes "\<And>x. g (h x) = x"
+ shows "foldl f (g s) xs = g (foldl (\<lambda>s x. h (f (g s) x)) s xs)"
+ by (rule sym, induct xs arbitrary: s) (simp_all add: assms)
+
+subsection {* Lifting *}
+
+datatype 'a fset = Fset "'a set"
+
+primrec member :: "'a fset \<Rightarrow> 'a set" where
+ "member (Fset A) = A"
+
+lemma Fset_member [simp]:
+ "Fset (member A) = A"
+ by (cases A) simp
+
+definition Set :: "'a list \<Rightarrow> 'a fset" where
+ "Set xs = Fset (set xs)"
+
+lemma member_Set [simp]:
+ "member (Set xs) = set xs"
+ by (simp add: Set_def)
+
+code_datatype Set
+
+
+subsection {* Basic operations *}
+
+definition is_empty :: "'a fset \<Rightarrow> bool" where
+ "is_empty A \<longleftrightarrow> List_Set.is_empty (member A)"
+
+lemma is_empty_Set [code]:
+ "is_empty (Set xs) \<longleftrightarrow> null xs"
+ by (simp add: is_empty_def is_empty_set)
+
+definition empty :: "'a fset" where
+ "empty = Fset {}"
+
+lemma empty_Set [code]:
+ "empty = Set []"
+ by (simp add: empty_def Set_def)
+
+definition insert :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
+ "insert x A = Fset (Set.insert x (member A))"
+
+lemma insert_Set [code]:
+ "insert x (Set xs) = Set (List_Set.insert x xs)"
+ by (simp add: insert_def Set_def insert_set)
+
+definition remove :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
+ "remove x A = Fset (List_Set.remove x (member A))"
+
+lemma remove_Set [code]:
+ "remove x (Set xs) = Set (remove_all x xs)"
+ by (simp add: remove_def Set_def remove_set)
+
+definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" where
+ "map f A = Fset (image f (member A))"
+
+lemma map_Set [code]:
+ "map f (Set xs) = Set (remdups (List.map f xs))"
+ by (simp add: map_def Set_def)
+
+definition project :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
+ "project P A = Fset (List_Set.project P (member A))"
+
+lemma project_Set [code]:
+ "project P (Set xs) = Set (filter P xs)"
+ by (simp add: project_def Set_def project_set)
+
+definition forall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> bool" where
+ "forall P A \<longleftrightarrow> Ball (member A) P"
+
+lemma forall_Set [code]:
+ "forall P (Set xs) \<longleftrightarrow> list_all P xs"
+ by (simp add: forall_def Set_def ball_set)
+
+definition exists :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> bool" where
+ "exists P A \<longleftrightarrow> Bex (member A) P"
+
+lemma exists_Set [code]:
+ "exists P (Set xs) \<longleftrightarrow> list_ex P xs"
+ by (simp add: exists_def Set_def bex_set)
+
+
+subsection {* Functorial operations *}
+
+definition union :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
+ "union A B = Fset (member A \<union> member B)"
+
+lemma union_insert [code]:
+ "union (Set xs) A = foldl (\<lambda>A x. insert x A) A xs"
+proof -
+ have "foldl (\<lambda>A x. Set.insert x A) (member A) xs =
+ member (foldl (\<lambda>A x. Fset (Set.insert x (member A))) A xs)"
+ by (rule foldl_apply_inv) simp
+ then show ?thesis by (simp add: union_def union_set insert_def)
+qed
+
+definition subtract :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
+ "subtract A B = Fset (member B - member A)"
+
+lemma subtract_remove [code]:
+ "subtract (Set xs) A = foldl (\<lambda>A x. remove x A) A xs"
+proof -
+ have "foldl (\<lambda>A x. List_Set.remove x A) (member A) xs =
+ member (foldl (\<lambda>A x. Fset (List_Set.remove x (member A))) A xs)"
+ by (rule foldl_apply_inv) simp
+ then show ?thesis by (simp add: subtract_def minus_set remove_def)
+qed
+
+
+subsection {* Derived operations *}
+
+lemma member_exists [code]:
+ "member A y \<longleftrightarrow> exists (\<lambda>x. y = x) A"
+ by (simp add: exists_def mem_def)
+
+definition subfset_eq :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" where
+ "subfset_eq A B \<longleftrightarrow> member A \<subseteq> member B"
+
+lemma subfset_eq_forall [code]:
+ "subfset_eq A B \<longleftrightarrow> forall (\<lambda>x. member B x) A"
+ by (simp add: subfset_eq_def subset_eq forall_def mem_def)
+
+definition subfset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" where
+ "subfset A B \<longleftrightarrow> member A \<subset> member B"
+
+lemma subfset_subfset_eq [code]:
+ "subfset A B \<longleftrightarrow> subfset_eq A B \<and> \<not> subfset_eq B A"
+ by (simp add: subfset_def subfset_eq_def subset)
+
+lemma eq_fset_subfset_eq [code]:
+ "eq_class.eq A B \<longleftrightarrow> subfset_eq A B \<and> subfset_eq B A"
+ by (cases A, cases B) (simp add: eq subfset_eq_def set_eq)
+
+definition inter :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
+ "inter A B = Fset (List_Set.project (member A) (member B))"
+
+lemma inter_project [code]:
+ "inter A B = project (member A) B"
+ by (simp add: inter_def project_def inter)
+
+
+subsection {* Misc operations *}
+
+lemma size_fset [code]:
+ "fset_size f A = 0"
+ "size A = 0"
+ by (cases A, simp) (cases A, simp)
+
+lemma fset_case_code [code]:
+ "fset_case f A = f (member A)"
+ by (cases A) simp
+
+lemma fset_rec_code [code]:
+ "fset_rec f A = f (member A)"
+ by (cases A) simp
+
+end
--- a/src/HOL/Library/Library.thy Thu Jun 25 16:48:43 2009 +0200
+++ b/src/HOL/Library/Library.thy Thu Jun 25 17:07:28 2009 +0200
@@ -10,6 +10,7 @@
Char_ord
Code_Char_chr
Code_Integer
+ Code_Set
Coinductive_List
Commutative_Ring
Continuity
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/List_Set.thy Thu Jun 25 17:07:28 2009 +0200
@@ -0,0 +1,163 @@
+
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* Relating (finite) sets and lists *}
+
+theory List_Set
+imports Main
+begin
+
+subsection {* Various additional list functions *}
+
+definition insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "insert x xs = (if x \<in> set xs then xs else x # xs)"
+
+definition remove_all :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "remove_all x xs = filter (Not o op = x) xs"
+
+
+subsection {* Various additional set functions *}
+
+definition is_empty :: "'a set \<Rightarrow> bool" where
+ "is_empty A \<longleftrightarrow> A = {}"
+
+definition remove :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+ "remove x A = A - {x}"
+
+lemma fun_left_comm_idem_remove:
+ "fun_left_comm_idem remove"
+proof -
+ have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: expand_fun_eq remove_def)
+ show ?thesis by (simp only: fun_left_comm_idem_remove rem)
+qed
+
+lemma minus_fold_remove:
+ assumes "finite A"
+ shows "B - A = fold remove B A"
+proof -
+ have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: expand_fun_eq remove_def)
+ show ?thesis by (simp only: rem assms minus_fold_remove)
+qed
+
+definition project :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+ "project P A = {a\<in>A. P a}"
+
+
+subsection {* Basic set operations *}
+
+lemma is_empty_set:
+ "is_empty (set xs) \<longleftrightarrow> null xs"
+ by (simp add: is_empty_def null_empty)
+
+lemma ball_set:
+ "(\<forall>x\<in>set xs. P x) \<longleftrightarrow> list_all P xs"
+ by (rule list_ball_code)
+
+lemma bex_set:
+ "(\<exists>x\<in>set xs. P x) \<longleftrightarrow> list_ex P xs"
+ by (rule list_bex_code)
+
+lemma empty_set:
+ "{} = set []"
+ by simp
+
+lemma insert_set:
+ "Set.insert x (set xs) = set (insert x xs)"
+ by (auto simp add: insert_def)
+
+lemma remove_set:
+ "remove x (set xs) = set (remove_all x xs)"
+ by (auto simp add: remove_def remove_all_def)
+
+lemma image_set:
+ "image f (set xs) = set (remdups (map f xs))"
+ by simp
+
+lemma project_set:
+ "project P (set xs) = set (filter P xs)"
+ by (auto simp add: project_def)
+
+
+subsection {* Functorial set operations *}
+
+lemma union_set:
+ "set xs \<union> A = foldl (\<lambda>A x. Set.insert x A) A xs"
+proof -
+ interpret fun_left_comm_idem Set.insert
+ by (fact fun_left_comm_idem_insert)
+ show ?thesis by (simp add: union_fold_insert fold_set)
+qed
+
+lemma minus_set:
+ "A - set xs = foldl (\<lambda>A x. remove x A) A xs"
+proof -
+ interpret fun_left_comm_idem remove
+ by (fact fun_left_comm_idem_remove)
+ show ?thesis
+ by (simp add: minus_fold_remove [of _ A] fold_set)
+qed
+
+lemma Inter_set:
+ "Inter (set (A # As)) = foldl (op \<inter>) A As"
+proof -
+ have "finite (set (A # As))" by simp
+ moreover have "fold (op \<inter>) UNIV (set (A # As)) = foldl (\<lambda>y x. x \<inter> y) UNIV (A # As)"
+ by (rule fun_left_comm_idem.fold_set, unfold_locales, auto)
+ ultimately have "Inter (set (A # As)) = foldl (op \<inter>) UNIV (A # As)"
+ by (simp only: Inter_fold_inter Int_commute)
+ then show ?thesis by simp
+qed
+
+lemma Union_set:
+ "Union (set As) = foldl (op \<union>) {} As"
+proof -
+ have "fold (op \<union>) {} (set As) = foldl (\<lambda>y x. x \<union> y) {} As"
+ by (rule fun_left_comm_idem.fold_set, unfold_locales, auto)
+ then show ?thesis
+ by (simp only: Union_fold_union finite_set Un_commute)
+qed
+
+lemma INTER_set:
+ "INTER (set (A # As)) f = foldl (\<lambda>B A. f A \<inter> B) (f A) As"
+proof -
+ have "finite (set (A # As))" by simp
+ moreover have "fold (\<lambda>A. op \<inter> (f A)) UNIV (set (A # As)) = foldl (\<lambda>B A. f A \<inter> B) UNIV (A # As)"
+ by (rule fun_left_comm_idem.fold_set, unfold_locales, auto)
+ ultimately have "INTER (set (A # As)) f = foldl (\<lambda>B A. f A \<inter> B) UNIV (A # As)"
+ by (simp only: INTER_fold_inter)
+ then show ?thesis by simp
+qed
+
+lemma UNION_set:
+ "UNION (set As) f = foldl (\<lambda>B A. f A \<union> B) {} As"
+proof -
+ have "fold (\<lambda>A. op \<union> (f A)) {} (set As) = foldl (\<lambda>B A. f A \<union> B) {} As"
+ by (rule fun_left_comm_idem.fold_set, unfold_locales, auto)
+ then show ?thesis
+ by (simp only: UNION_fold_union finite_set)
+qed
+
+
+subsection {* Derived set operations *}
+
+lemma member:
+ "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. a = x)"
+ by simp
+
+lemma subset_eq:
+ "A \<subseteq> B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
+ by (fact subset_eq)
+
+lemma subset:
+ "A \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> \<not> B \<subseteq> A"
+ by (fact less_le_not_le)
+
+lemma set_eq:
+ "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
+ by (fact eq_iff)
+
+lemma inter:
+ "A \<inter> B = project (\<lambda>x. x \<in> A) B"
+ by (auto simp add: project_def)
+
+end
\ No newline at end of file
--- a/src/HOL/ex/Codegenerator_Candidates.thy Thu Jun 25 16:48:43 2009 +0200
+++ b/src/HOL/ex/Codegenerator_Candidates.thy Thu Jun 25 17:07:28 2009 +0200
@@ -8,6 +8,7 @@
Complex_Main
AssocList
Binomial
+ Code_Set
Commutative_Ring
Enum
List_Prefix