# HG changeset patch # User Andreas Lochbihler # Date 1360854088 -3600 # Node ID 0dac0158b8d46a59c4e474526df5cc1107aac1ef # Parent 3e913a575dc677215aafc07aa63b44d0e7b68e09 implement code generation for finite, card, op = and op <= for sets always via finite_UNIV and card_UNIV, as fragile rewrites based on sorts are hard to find and debug diff -r 3e913a575dc6 -r 0dac0158b8d4 src/HOL/Codegenerator_Test/RBT_Set_Test.thy --- a/src/HOL/Codegenerator_Test/RBT_Set_Test.thy Thu Feb 14 12:24:56 2013 +0100 +++ b/src/HOL/Codegenerator_Test/RBT_Set_Test.thy Thu Feb 14 16:01:28 2013 +0100 @@ -22,23 +22,13 @@ lemma [code, code del]: "pred_of_set = pred_of_set" .. - -lemma [code, code del]: - "Cardinality.card' = Cardinality.card'" .. - -lemma [code, code del]: - "Cardinality.finite' = Cardinality.finite'" .. - -lemma [code, code del]: - "Cardinality.subset' = Cardinality.subset'" .. - -lemma [code, code del]: - "Cardinality.eq_set = Cardinality.eq_set" .. - - lemma [code, code del]: "acc = acc" .. +lemmas [code del] = + finite_set_code finite_coset_code + equal_set_code + (* If the code generation ends with an exception with the following message: '"List.set" is not a constructor, on left hand side of equation: ...', diff -r 3e913a575dc6 -r 0dac0158b8d4 src/HOL/Library/Cardinality.thy --- a/src/HOL/Library/Cardinality.thy Thu Feb 14 12:24:56 2013 +0100 +++ b/src/HOL/Library/Cardinality.thy Thu Feb 14 16:01:28 2013 +0100 @@ -199,7 +199,7 @@ subsection {* A type class for computing the cardinality of types *} definition is_list_UNIV :: "'a list \ bool" -where [code del]: "is_list_UNIV xs = (let c = CARD('a) in if c = 0 then False else size (remdups xs) = c)" +where "is_list_UNIV xs = (let c = CARD('a) in if c = 0 then False else size (remdups xs) = c)" lemma is_list_UNIV_iff: "is_list_UNIV xs \ set xs = UNIV" by(auto simp add: is_list_UNIV_def Let_def card_eq_0_iff List.card_set[symmetric] @@ -211,15 +211,6 @@ fixes card_UNIV :: "'a card_UNIV" assumes card_UNIV: "card_UNIV = Phantom('a) CARD('a)" -lemma card_UNIV_code [code_unfold]: - "CARD('a :: card_UNIV) = of_phantom (card_UNIV :: 'a card_UNIV)" -by(simp add: card_UNIV) - -lemma is_list_UNIV_code [code]: - "is_list_UNIV (xs :: 'a list) = - (let c = CARD('a :: card_UNIV) in if c = 0 then False else size (remdups xs) = c)" -by(rule is_list_UNIV_def) - subsection {* Instantiations for @{text "card_UNIV"} *} instantiation nat :: card_UNIV begin @@ -396,80 +387,66 @@ subsection {* Code setup for sets *} +text {* + Implement operations @{term "finite"}, @{term "card"}, @{term "op \"}, and @{term "op ="} + for sets using @{term "finite_UNIV"} and @{term "card_UNIV"}. +*} + lemma card_Compl: "finite A \ card (- A) = card (UNIV :: 'a set) - card (A :: 'a set)" by (metis Compl_eq_Diff_UNIV card_Diff_subset top_greatest) -context fixes xs :: "'a :: card_UNIV list" -begin +lemma card_coset_code [code]: + fixes xs :: "'a :: card_UNIV list" + shows "card (List.coset xs) = of_phantom (card_UNIV :: 'a card_UNIV) - length (remdups xs)" +by(simp add: List.card_set card_Compl card_UNIV) -definition card' :: "'a set \ nat" -where [simp, code del, code_abbrev]: "card' = card" - -lemma card'_code [code]: - "card' (set xs) = length (remdups xs)" - "card' (List.coset xs) = of_phantom (card_UNIV :: 'a card_UNIV) - length (remdups xs)" -by(simp_all add: List.card_set card_Compl card_UNIV) +lemma [code, code del]: "finite = finite" .. -lemma card'_UNIV [code_unfold]: - "card' (UNIV :: 'a :: card_UNIV set) = of_phantom (card_UNIV :: 'a card_UNIV)" -by(simp add: card_UNIV) - -definition finite' :: "'a set \ bool" -where [simp, code del, code_abbrev]: "finite' = finite" - -lemma finite'_code [code]: - "finite' (set xs) \ True" - "finite' (List.coset xs) \ of_phantom (finite_UNIV :: 'a finite_UNIV)" +lemma [code]: + fixes xs :: "'a :: card_UNIV list" + shows finite_set_code: + "finite (set xs) = True" + and finite_coset_code: + "finite (List.coset xs) \ of_phantom (finite_UNIV :: 'a finite_UNIV)" by(simp_all add: card_gt_0_iff finite_UNIV) -definition subset' :: "'a set \ 'a set \ bool" -where [simp, code del, code_abbrev]: "subset' = op \" - -lemma subset'_code [code]: - "subset' A (List.coset ys) \ (\y \ set ys. y \ A)" - "subset' (set ys) B \ (\y \ set ys. y \ B)" - "subset' (List.coset xs) (set ys) \ (let n = CARD('a) in n > 0 \ card(set (xs @ ys)) = n)" +lemma coset_subset_code [code]: + fixes xs :: "'a list" shows + "List.coset xs \ set ys \ (let n = CARD('a) in n > 0 \ card (set (xs @ ys)) = n)" by(auto simp add: Let_def card_gt_0_iff dest: card_eq_UNIV_imp_eq_UNIV intro: arg_cong[where f=card]) (metis finite_compl finite_set rev_finite_subset) -definition eq_set :: "'a set \ 'a set \ bool" -where [simp, code del, code_abbrev]: "eq_set = op =" - -lemma eq_set_code [code]: - fixes ys +lemma equal_set_code [code]: + fixes xs ys :: "'a :: equal list" defines "rhs \ let n = CARD('a) in if n = 0 then False else let xs' = remdups xs; ys' = remdups ys in length xs' + length ys' = n \ (\x \ set xs'. x \ set ys') \ (\y \ set ys'. y \ set xs')" - shows "eq_set (List.coset xs) (set ys) \ rhs" (is ?thesis1) - and "eq_set (set ys) (List.coset xs) \ rhs" (is ?thesis2) - and "eq_set (set xs) (set ys) \ (\x \ set xs. x \ set ys) \ (\y \ set ys. y \ set xs)" (is ?thesis3) - and "eq_set (List.coset xs) (List.coset ys) \ (\x \ set xs. x \ set ys) \ (\y \ set ys. y \ set xs)" (is ?thesis4) + shows "equal_class.equal (List.coset xs) (set ys) \ rhs" (is ?thesis1) + and "equal_class.equal (set ys) (List.coset xs) \ rhs" (is ?thesis2) + and "equal_class.equal (set xs) (set ys) \ (\x \ set xs. x \ set ys) \ (\y \ set ys. y \ set xs)" (is ?thesis3) + and "equal_class.equal (List.coset xs) (List.coset ys) \ (\x \ set xs. x \ set ys) \ (\y \ set ys. y \ set xs)" (is ?thesis4) proof - show ?thesis1 (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs - by(auto simp add: rhs_def Let_def List.card_set[symmetric] card_Un_Int[where A="set xs" and B="- set xs"] card_UNIV Compl_partition card_gt_0_iff dest: sym)(metis finite_compl finite_set) + by(auto simp add: equal_eq rhs_def Let_def List.card_set[symmetric] card_Un_Int[where A="set xs" and B="- set xs"] card_UNIV Compl_partition card_gt_0_iff dest: sym)(metis finite_compl finite_set) next assume ?rhs moreover have "\ \y\set xs. y \ set ys; \x\set ys. x \ set xs \ \ set xs \ set ys = {}" by blast ultimately show ?lhs - by(auto simp add: rhs_def Let_def List.card_set[symmetric] card_UNIV card_gt_0_iff card_Un_Int[where A="set xs" and B="set ys"] dest: card_eq_UNIV_imp_eq_UNIV split: split_if_asm) + by(auto simp add: equal_eq rhs_def Let_def List.card_set[symmetric] card_UNIV card_gt_0_iff card_Un_Int[where A="set xs" and B="set ys"] dest: card_eq_UNIV_imp_eq_UNIV split: split_if_asm) qed - thus ?thesis2 unfolding eq_set_def by blast - show ?thesis3 ?thesis4 unfolding eq_set_def List.coset_def by blast+ + thus ?thesis2 unfolding equal_eq by blast + show ?thesis3 ?thesis4 unfolding equal_eq List.coset_def by blast+ qed -end - notepad begin (* test code setup *) have "List.coset [True] = set [False] \ List.coset [] \ List.set [True, False] \ finite (List.coset [True])" by eval end -hide_const (open) card' finite' subset' eq_set - end