# HG changeset patch # User bulwahn # Date 1307438656 -7200 # Node ID 3c58977e0911b0f96cb25280ce133b57d0f2460e # Parent 93b1183e43e548107c90af993c28ab187dede129# Parent 9d717505595f24b20bb0fe0c4e4799ba23297c9a merged; manually merged IsaMakefile diff -r 9d717505595f -r 3c58977e0911 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Tue Jun 07 11:05:09 2011 +0200 +++ b/src/HOL/IsaMakefile Tue Jun 07 11:24:16 2011 +0200 @@ -453,7 +453,7 @@ Library/Indicator_Function.thy Library/Infinite_Set.thy \ Library/Inner_Product.thy Library/Kleene_Algebra.thy \ Library/LaTeXsugar.thy Library/Lattice_Algebras.thy \ - Library/Lattice_Syntax.thy Library/Library.thy \ + Library/Lattice_Syntax.thy Library/Library.thy Library/List_Cset.thy \ Library/List_Prefix.thy Library/List_lexord.thy Library/Mapping.thy \ Library/Monad_Syntax.thy Library/More_List.thy Library/More_Set.thy \ Library/Multiset.thy Library/Nat_Bijection.thy \ @@ -1047,7 +1047,7 @@ $(LOG)/HOL-ex.gz: $(OUT)/HOL Decision_Procs/Commutative_Ring.thy \ Number_Theory/Primes.thy ex/Abstract_NAT.thy ex/Antiquote.thy \ ex/ATP_Export.thy ex/Arith_Examples.thy ex/Arithmetic_Series_Complex.thy \ - ex/BT.thy ex/BinEx.thy ex/Binary.thy ex/Birthday_Paradoxon.thy \ + ex/BT.thy ex/BinEx.thy ex/Binary.thy ex/Birthday_Paradox.thy \ ex/CASC_Setup.thy ex/CTL.thy ex/Case_Product.thy \ ex/Chinese.thy ex/Classical.thy ex/CodegenSML_Test.thy \ ex/Coercion_Examples.thy ex/Coherent.thy ex/Dedekind_Real.thy \ diff -r 9d717505595f -r 3c58977e0911 src/HOL/Library/Cset.thy --- a/src/HOL/Library/Cset.thy Tue Jun 07 11:05:09 2011 +0200 +++ b/src/HOL/Library/Cset.thy Tue Jun 07 11:24:16 2011 +0200 @@ -35,66 +35,6 @@ by (simp add: Cset.set_eq_iff) hide_fact (open) set_eqI -declare mem_def [simp] - -definition set :: "'a list \ 'a Cset.set" where - "set xs = Set (List.set xs)" -hide_const (open) set - -lemma member_set [simp]: - "member (Cset.set xs) = set xs" - by (simp add: set_def) -hide_fact (open) member_set - -definition coset :: "'a list \ 'a Cset.set" where - "coset xs = Set (- set xs)" -hide_const (open) coset - -lemma member_coset [simp]: - "member (Cset.coset xs) = - set xs" - by (simp add: coset_def) -hide_fact (open) member_coset - -code_datatype Cset.set Cset.coset - -lemma member_code [code]: - "member (Cset.set xs) = List.member xs" - "member (Cset.coset xs) = Not \ List.member xs" - by (simp_all add: fun_eq_iff member_def fun_Compl_def bool_Compl_def) - -lemma member_image_UNIV [simp]: - "member ` UNIV = UNIV" -proof - - have "\A \ 'a set. \B \ 'a Cset.set. A = member B" - proof - fix A :: "'a set" - show "A = member (Set A)" by simp - qed - then show ?thesis by (simp add: image_def) -qed - -definition (in term_syntax) - setify :: "'a\typerep list \ (unit \ Code_Evaluation.term) - \ 'a Cset.set \ (unit \ Code_Evaluation.term)" where - [code_unfold]: "setify xs = Code_Evaluation.valtermify Cset.set {\} xs" - -notation fcomp (infixl "\>" 60) -notation scomp (infixl "\\" 60) - -instantiation Cset.set :: (random) random -begin - -definition - "Quickcheck.random i = Quickcheck.random i \\ (\xs. Pair (setify xs))" - -instance .. - -end - -no_notation fcomp (infixl "\>" 60) -no_notation scomp (infixl "\\" 60) - - subsection {* Lattice instantiation *} instantiation Cset.set :: (type) boolean_algebra @@ -149,185 +89,39 @@ definition is_empty :: "'a Cset.set \ bool" where [simp]: "is_empty A \ More_Set.is_empty (member A)" -lemma is_empty_set [code]: - "is_empty (Cset.set xs) \ List.null xs" - by (simp add: is_empty_set) -hide_fact (open) is_empty_set - -lemma empty_set [code]: - "bot = Cset.set []" - by (simp add: set_def) -hide_fact (open) empty_set - -lemma UNIV_set [code]: - "top = Cset.coset []" - by (simp add: coset_def) -hide_fact (open) UNIV_set - definition insert :: "'a \ 'a Cset.set \ 'a Cset.set" where [simp]: "insert x A = Set (Set.insert x (member A))" -lemma insert_set [code]: - "insert x (Cset.set xs) = Cset.set (List.insert x xs)" - "insert x (Cset.coset xs) = Cset.coset (removeAll x xs)" - by (simp_all add: set_def coset_def) - definition remove :: "'a \ 'a Cset.set \ 'a Cset.set" where [simp]: "remove x A = Set (More_Set.remove x (member A))" -lemma remove_set [code]: - "remove x (Cset.set xs) = Cset.set (removeAll x xs)" - "remove x (Cset.coset xs) = Cset.coset (List.insert x xs)" - by (simp_all add: set_def coset_def remove_set_compl) - (simp add: More_Set.remove_def) - definition map :: "('a \ 'b) \ 'a Cset.set \ 'b Cset.set" where [simp]: "map f A = Set (image f (member A))" -lemma map_set [code]: - "map f (Cset.set xs) = Cset.set (remdups (List.map f xs))" - by (simp add: set_def) - enriched_type map: map by (simp_all add: fun_eq_iff image_compose) definition filter :: "('a \ bool) \ 'a Cset.set \ 'a Cset.set" where [simp]: "filter P A = Set (More_Set.project P (member A))" -lemma filter_set [code]: - "filter P (Cset.set xs) = Cset.set (List.filter P xs)" - by (simp add: set_def project_set) - definition forall :: "('a \ bool) \ 'a Cset.set \ bool" where [simp]: "forall P A \ Ball (member A) P" -lemma forall_set [code]: - "forall P (Cset.set xs) \ list_all P xs" - by (simp add: set_def list_all_iff) - definition exists :: "('a \ bool) \ 'a Cset.set \ bool" where [simp]: "exists P A \ Bex (member A) P" -lemma exists_set [code]: - "exists P (Cset.set xs) \ list_ex P xs" - by (simp add: set_def list_ex_iff) - definition card :: "'a Cset.set \ nat" where [simp]: "card A = Finite_Set.card (member A)" - -lemma card_set [code]: - "card (Cset.set xs) = length (remdups xs)" -proof - - have "Finite_Set.card (set (remdups xs)) = length (remdups xs)" - by (rule distinct_card) simp - then show ?thesis by (simp add: set_def) -qed - -lemma compl_set [simp, code]: - "- Cset.set xs = Cset.coset xs" - by (simp add: set_def coset_def) - -lemma compl_coset [simp, code]: - "- Cset.coset xs = Cset.set xs" - by (simp add: set_def coset_def) - - -subsection {* Derived operations *} - -lemma subset_eq_forall [code]: - "A \ B \ forall (member B) A" - by (simp add: subset_eq) - -lemma subset_subset_eq [code]: - "A < B \ A \ B \ \ B \ (A :: 'a Cset.set)" - by (fact less_le_not_le) - -instantiation Cset.set :: (type) equal -begin - -definition [code]: - "HOL.equal A B \ A \ B \ B \ (A :: 'a Cset.set)" - -instance proof -qed (simp add: equal_set_def set_eq [symmetric] Cset.set_eq_iff) - -end - -lemma [code nbe]: - "HOL.equal (A :: 'a Cset.set) A \ True" - by (fact equal_refl) - - -subsection {* Functorial operations *} - -lemma inter_project [code]: - "inf A (Cset.set xs) = Cset.set (List.filter (member A) xs)" - "inf A (Cset.coset xs) = foldr remove xs A" -proof - - show "inf A (Cset.set xs) = Cset.set (List.filter (member A) xs)" - by (simp add: inter project_def set_def) - have *: "\x::'a. remove = (\x. Set \ More_Set.remove x \ member)" - by (simp add: fun_eq_iff) - have "member \ fold (\x. Set \ More_Set.remove x \ member) xs = - fold More_Set.remove xs \ member" - by (rule fold_commute) (simp add: fun_eq_iff) - then have "fold More_Set.remove xs (member A) = - member (fold (\x. Set \ More_Set.remove x \ member) xs A)" - by (simp add: fun_eq_iff) - then have "inf A (Cset.coset xs) = fold remove xs A" - by (simp add: Diff_eq [symmetric] minus_set *) - moreover have "\x y :: 'a. Cset.remove y \ Cset.remove x = Cset.remove x \ Cset.remove y" - by (auto simp add: More_Set.remove_def * intro: ext) - ultimately show "inf A (Cset.coset xs) = foldr remove xs A" - by (simp add: foldr_fold) -qed - -lemma subtract_remove [code]: - "A - Cset.set xs = foldr remove xs A" - "A - Cset.coset xs = Cset.set (List.filter (member A) xs)" - by (simp_all only: diff_eq compl_set compl_coset inter_project) - -lemma union_insert [code]: - "sup (Cset.set xs) A = foldr insert xs A" - "sup (Cset.coset xs) A = Cset.coset (List.filter (Not \ member A) xs)" -proof - - have *: "\x::'a. insert = (\x. Set \ Set.insert x \ member)" - by (simp add: fun_eq_iff) - have "member \ fold (\x. Set \ Set.insert x \ member) xs = - fold Set.insert xs \ member" - by (rule fold_commute) (simp add: fun_eq_iff) - then have "fold Set.insert xs (member A) = - member (fold (\x. Set \ Set.insert x \ member) xs A)" - by (simp add: fun_eq_iff) - then have "sup (Cset.set xs) A = fold insert xs A" - by (simp add: union_set *) - moreover have "\x y :: 'a. Cset.insert y \ Cset.insert x = Cset.insert x \ Cset.insert y" - by (auto simp add: * intro: ext) - ultimately show "sup (Cset.set xs) A = foldr insert xs A" - by (simp add: foldr_fold) - show "sup (Cset.coset xs) A = Cset.coset (List.filter (Not \ member A) xs)" - by (auto simp add: coset_def) -qed - + context complete_lattice begin definition Infimum :: "'a Cset.set \ 'a" where [simp]: "Infimum A = Inf (member A)" -lemma Infimum_inf [code]: - "Infimum (Cset.set As) = foldr inf As top" - "Infimum (Cset.coset []) = bot" - by (simp_all add: Inf_set_foldr Inf_UNIV) - definition Supremum :: "'a Cset.set \ 'a" where [simp]: "Supremum A = Sup (member A)" -lemma Supremum_sup [code]: - "Supremum (Cset.set As) = foldr sup As bot" - "Supremum (Cset.coset []) = top" - by (simp_all add: Sup_set_foldr Sup_UNIV) - end @@ -351,7 +145,7 @@ declare mem_def [simp del] -hide_const (open) setify is_empty insert remove map filter forall exists card +hide_const (open) is_empty insert remove map filter forall exists card Inter Union end diff -r 9d717505595f -r 3c58977e0911 src/HOL/Library/Dlist_Cset.thy --- a/src/HOL/Library/Dlist_Cset.thy Tue Jun 07 11:05:09 2011 +0200 +++ b/src/HOL/Library/Dlist_Cset.thy Tue Jun 07 11:24:16 2011 +0200 @@ -3,21 +3,21 @@ header {* Canonical implementation of sets by distinct lists *} theory Dlist_Cset -imports Dlist Cset +imports Dlist List_Cset begin definition Set :: "'a dlist \ 'a Cset.set" where - "Set dxs = Cset.set (list_of_dlist dxs)" + "Set dxs = List_Cset.set (list_of_dlist dxs)" definition Coset :: "'a dlist \ 'a Cset.set" where - "Coset dxs = Cset.coset (list_of_dlist dxs)" + "Coset dxs = List_Cset.coset (list_of_dlist dxs)" code_datatype Set Coset declare member_code [code del] -declare Cset.is_empty_set [code del] -declare Cset.empty_set [code del] -declare Cset.UNIV_set [code del] +declare List_Cset.is_empty_set [code del] +declare List_Cset.empty_set [code del] +declare List_Cset.UNIV_set [code del] declare insert_set [code del] declare remove_set [code del] declare compl_set [code del] @@ -50,11 +50,11 @@ by (simp add: Coset_def member_set not_set_compl) lemma Set_dlist_of_list [code]: - "Cset.set xs = Set (dlist_of_list xs)" + "List_Cset.set xs = Set (dlist_of_list xs)" by (rule Cset.set_eqI) simp lemma Coset_dlist_of_list [code]: - "Cset.coset xs = Coset (dlist_of_list xs)" + "List_Cset.coset xs = Coset (dlist_of_list xs)" by (rule Cset.set_eqI) simp lemma is_empty_Set [code]: diff -r 9d717505595f -r 3c58977e0911 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Tue Jun 07 11:05:09 2011 +0200 +++ b/src/HOL/Library/Library.thy Tue Jun 07 11:24:16 2011 +0200 @@ -29,6 +29,7 @@ Lattice_Algebras Lattice_Syntax ListVector + List_Cset Kleene_Algebra Mapping Monad_Syntax diff -r 9d717505595f -r 3c58977e0911 src/HOL/Library/List_Cset.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/List_Cset.thy Tue Jun 07 11:24:16 2011 +0200 @@ -0,0 +1,222 @@ + +(* Author: Florian Haftmann, TU Muenchen *) + +header {* implementation of Cset.sets based on lists *} + +theory List_Cset +imports Cset +begin + +declare mem_def [simp] + +definition set :: "'a list \ 'a Cset.set" where + "set xs = Set (List.set xs)" +hide_const (open) set + +lemma member_set [simp]: + "member (List_Cset.set xs) = set xs" + by (simp add: set_def) +hide_fact (open) member_set + +definition coset :: "'a list \ 'a Cset.set" where + "coset xs = Set (- set xs)" +hide_const (open) coset + +lemma member_coset [simp]: + "member (List_Cset.coset xs) = - set xs" + by (simp add: coset_def) +hide_fact (open) member_coset + +code_datatype List_Cset.set List_Cset.coset + +lemma member_code [code]: + "member (List_Cset.set xs) = List.member xs" + "member (List_Cset.coset xs) = Not \ List.member xs" + by (simp_all add: fun_eq_iff member_def fun_Compl_def bool_Compl_def) + +lemma member_image_UNIV [simp]: + "member ` UNIV = UNIV" +proof - + have "\A \ 'a set. \B \ 'a Cset.set. A = member B" + proof + fix A :: "'a set" + show "A = member (Set A)" by simp + qed + then show ?thesis by (simp add: image_def) +qed + +definition (in term_syntax) + setify :: "'a\typerep list \ (unit \ Code_Evaluation.term) + \ 'a Cset.set \ (unit \ Code_Evaluation.term)" where + [code_unfold]: "setify xs = Code_Evaluation.valtermify List_Cset.set {\} xs" + +notation fcomp (infixl "\>" 60) +notation scomp (infixl "\\" 60) + +instantiation Cset.set :: (random) random +begin + +definition + "Quickcheck.random i = Quickcheck.random i \\ (\xs. Pair (setify xs))" + +instance .. + +end + +no_notation fcomp (infixl "\>" 60) +no_notation scomp (infixl "\\" 60) + +subsection {* Basic operations *} + +lemma is_empty_set [code]: + "Cset.is_empty (List_Cset.set xs) \ List.null xs" + by (simp add: is_empty_set null_def) +hide_fact (open) is_empty_set + +lemma empty_set [code]: + "bot = List_Cset.set []" + by (simp add: set_def) +hide_fact (open) empty_set + +lemma UNIV_set [code]: + "top = List_Cset.coset []" + by (simp add: coset_def) +hide_fact (open) UNIV_set + +lemma remove_set [code]: + "Cset.remove x (List_Cset.set xs) = List_Cset.set (removeAll x xs)" + "Cset.remove x (List_Cset.coset xs) = List_Cset.coset (List.insert x xs)" +by (simp_all add: set_def coset_def) + (metis List.set_insert More_Set.remove_def remove_set_compl) + +lemma insert_set [code]: + "Cset.insert x (List_Cset.set xs) = List_Cset.set (List.insert x xs)" + "Cset.insert x (List_Cset.coset xs) = List_Cset.coset (removeAll x xs)" + by (simp_all add: set_def coset_def) + +lemma map_set [code]: + "Cset.map f (List_Cset.set xs) = List_Cset.set (remdups (List.map f xs))" + by (simp add: set_def) + +lemma filter_set [code]: + "Cset.filter P (List_Cset.set xs) = List_Cset.set (List.filter P xs)" + by (simp add: set_def project_set) + +lemma forall_set [code]: + "Cset.forall P (List_Cset.set xs) \ list_all P xs" + by (simp add: set_def list_all_iff) + +lemma exists_set [code]: + "Cset.exists P (List_Cset.set xs) \ list_ex P xs" + by (simp add: set_def list_ex_iff) + +lemma card_set [code]: + "Cset.card (List_Cset.set xs) = length (remdups xs)" +proof - + have "Finite_Set.card (set (remdups xs)) = length (remdups xs)" + by (rule distinct_card) simp + then show ?thesis by (simp add: set_def) +qed + +lemma compl_set [simp, code]: + "- List_Cset.set xs = List_Cset.coset xs" + by (simp add: set_def coset_def) + +lemma compl_coset [simp, code]: + "- List_Cset.coset xs = List_Cset.set xs" + by (simp add: set_def coset_def) + +context complete_lattice +begin + +lemma Infimum_inf [code]: + "Infimum (List_Cset.set As) = foldr inf As top" + "Infimum (List_Cset.coset []) = bot" + by (simp_all add: Inf_set_foldr Inf_UNIV) + +lemma Supremum_sup [code]: + "Supremum (List_Cset.set As) = foldr sup As bot" + "Supremum (List_Cset.coset []) = top" + by (simp_all add: Sup_set_foldr Sup_UNIV) + +end + + +subsection {* Derived operations *} + +lemma subset_eq_forall [code]: + "A \ B \ Cset.forall (member B) A" + by (simp add: subset_eq) + +lemma subset_subset_eq [code]: + "A < B \ A \ B \ \ B \ (A :: 'a Cset.set)" + by (fact less_le_not_le) + +instantiation Cset.set :: (type) equal +begin + +definition [code]: + "HOL.equal A B \ A \ B \ B \ (A :: 'a Cset.set)" + +instance proof +qed (simp add: equal_set_def set_eq [symmetric] Cset.set_eq_iff) + +end + +lemma [code nbe]: + "HOL.equal (A :: 'a Cset.set) A \ True" + by (fact equal_refl) + + +subsection {* Functorial operations *} + +lemma inter_project [code]: + "inf A (List_Cset.set xs) = List_Cset.set (List.filter (Cset.member A) xs)" + "inf A (List_Cset.coset xs) = foldr Cset.remove xs A" +proof - + show "inf A (List_Cset.set xs) = List_Cset.set (List.filter (member A) xs)" + by (simp add: inter project_def set_def) + have *: "\x::'a. Cset.remove = (\x. Set \ More_Set.remove x \ member)" + by (simp add: fun_eq_iff More_Set.remove_def) + have "member \ fold (\x. Set \ More_Set.remove x \ member) xs = + fold More_Set.remove xs \ member" + by (rule fold_commute) (simp add: fun_eq_iff) + then have "fold More_Set.remove xs (member A) = + member (fold (\x. Set \ More_Set.remove x \ member) xs A)" + by (simp add: fun_eq_iff) + then have "inf A (List_Cset.coset xs) = fold Cset.remove xs A" + by (simp add: Diff_eq [symmetric] minus_set *) + moreover have "\x y :: 'a. Cset.remove y \ Cset.remove x = Cset.remove x \ Cset.remove y" + by (auto simp add: More_Set.remove_def * intro: ext) + ultimately show "inf A (List_Cset.coset xs) = foldr Cset.remove xs A" + by (simp add: foldr_fold) +qed + +lemma subtract_remove [code]: + "A - List_Cset.set xs = foldr Cset.remove xs A" + "A - List_Cset.coset xs = List_Cset.set (List.filter (member A) xs)" + by (simp_all only: diff_eq compl_set compl_coset inter_project) + +lemma union_insert [code]: + "sup (List_Cset.set xs) A = foldr Cset.insert xs A" + "sup (List_Cset.coset xs) A = List_Cset.coset (List.filter (Not \ member A) xs)" +proof - + have *: "\x::'a. Cset.insert = (\x. Set \ Set.insert x \ member)" + by (simp add: fun_eq_iff) + have "member \ fold (\x. Set \ Set.insert x \ member) xs = + fold Set.insert xs \ member" + by (rule fold_commute) (simp add: fun_eq_iff) + then have "fold Set.insert xs (member A) = + member (fold (\x. Set \ Set.insert x \ member) xs A)" + by (simp add: fun_eq_iff) + then have "sup (List_Cset.set xs) A = fold Cset.insert xs A" + by (simp add: union_set *) + moreover have "\x y :: 'a. Cset.insert y \ Cset.insert x = Cset.insert x \ Cset.insert y" + by (auto simp add: * intro: ext) + ultimately show "sup (List_Cset.set xs) A = foldr Cset.insert xs A" + by (simp add: foldr_fold) + show "sup (List_Cset.coset xs) A = List_Cset.coset (List.filter (Not \ member A) xs)" + by (auto simp add: coset_def) +qed + +end \ No newline at end of file diff -r 9d717505595f -r 3c58977e0911 src/HOL/Library/Quickcheck_Narrowing.thy --- a/src/HOL/Library/Quickcheck_Narrowing.thy Tue Jun 07 11:05:09 2011 +0200 +++ b/src/HOL/Library/Quickcheck_Narrowing.thy Tue Jun 07 11:24:16 2011 +0200 @@ -5,6 +5,8 @@ theory Quickcheck_Narrowing imports Main "~~/src/HOL/Library/Code_Char" uses + ("~~/src/HOL/Tools/Quickcheck/PNF_Narrowing_Engine.hs") + ("~~/src/HOL/Tools/Quickcheck/Narrowing_Engine.hs") ("~~/src/HOL/Tools/Quickcheck/narrowing_generators.ML") begin @@ -454,6 +456,17 @@ end +datatype property = Universal narrowing_type "(narrowing_term => property)" "narrowing_term => Code_Evaluation.term" | Existential narrowing_type "(narrowing_term => property)" "narrowing_term => Code_Evaluation.term" | Property bool + +(* FIXME: hard-wired maximal depth of 100 here *) +fun exists :: "('a :: {narrowing, partial_term_of} => property) => property" +where + "exists f = (case narrowing (100 :: code_int) of C ty cs => Existential ty (\ t. f (conv cs t)) (partial_term_of (TYPE('a))))" + +fun "all" :: "('a :: {narrowing, partial_term_of} => property) => property" +where + "all f = (case narrowing (100 :: code_int) of C ty cs => Universal ty (\t. f (conv cs t)) (partial_term_of (TYPE('a))))" + subsubsection {* class @{text is_testable} *} text {* The class @{text is_testable} ensures that all necessary type instances are generated. *} @@ -492,13 +505,15 @@ hide_const (open) Constant eval_cfun subsubsection {* Setting up the counterexample generator *} - + +setup {* Thy_Load.provide_file (Path.explode ("~~/src/HOL/Tools/Quickcheck/PNF_Narrowing_Engine.hs")) *} +setup {* Thy_Load.provide_file (Path.explode ("~~/src/HOL/Tools/Quickcheck/Narrowing_Engine.hs")) *} use "~~/src/HOL/Tools/Quickcheck/narrowing_generators.ML" setup {* Narrowing_Generators.setup *} hide_type (open) code_int narrowing_type narrowing_term cons hide_const (open) int_of of_int nth error toEnum map_index split_At empty - C cons conv nonEmpty "apply" sum cons1 cons2 ensure_testable + C cons conv nonEmpty "apply" sum cons1 cons2 ensure_testable all exists end \ No newline at end of file diff -r 9d717505595f -r 3c58977e0911 src/HOL/Tools/Quickcheck/narrowing_generators.ML --- a/src/HOL/Tools/Quickcheck/narrowing_generators.ML Tue Jun 07 11:05:09 2011 +0200 +++ b/src/HOL/Tools/Quickcheck/narrowing_generators.ML Tue Jun 07 11:24:16 2011 +0200 @@ -6,10 +6,15 @@ signature NARROWING_GENERATORS = sig - val test_term: Proof.context -> bool * bool -> term * term list -> Quickcheck.result - val put_counterexample: (unit -> term list option) -> Proof.context -> Proof.context + val allow_existentials : bool Config.T val finite_functions : bool Config.T val overlord : bool Config.T + val test_term: Proof.context -> bool * bool -> term * term list -> Quickcheck.result + datatype counterexample = Universal_Counterexample of (term * counterexample) + | Existential_Counterexample of (term * counterexample) list + | Empty_Assignment + val put_counterexample: (unit -> term list option) -> Proof.context -> Proof.context + val put_existential_counterexample : (unit -> counterexample option) -> Proof.context -> Proof.context val setup: theory -> theory end; @@ -18,6 +23,7 @@ (* configurations *) +val allow_existentials = Attrib.setup_config_bool @{binding quickcheck_allow_existentials} (K true) val finite_functions = Attrib.setup_config_bool @{binding quickcheck_finite_functions} (K true) val overlord = Attrib.setup_config_bool @{binding quickcheck_narrowing_overlord} (K false) @@ -165,18 +171,25 @@ in eqs end - + +fun contains_recursive_type_under_function_types xs = + exists (fn (_, (_, _, cs)) => cs |> exists (snd #> exists (fn dT => + (case Datatype_Aux.strip_dtyp dT of (_ :: _, Datatype.DtRec _) => true | _ => false)))) xs + fun instantiate_narrowing_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy = let val _ = Datatype_Aux.message config "Creating narrowing generators ..."; val narrowingsN = map (prefix (narrowingN ^ "_")) (names @ auxnames); in - thy - |> Class.instantiation (tycos, vs, @{sort narrowing}) - |> Quickcheck_Common.define_functions - (fn narrowings => mk_equations descr vs tycos narrowings (Ts, Us), NONE) - prfx [] narrowingsN (map narrowingT (Ts @ Us)) - |> Class.prove_instantiation_exit (K (Class.intro_classes_tac [])) + if not (contains_recursive_type_under_function_types descr) then + thy + |> Class.instantiation (tycos, vs, @{sort narrowing}) + |> Quickcheck_Common.define_functions + (fn narrowings => mk_equations descr vs tycos narrowings (Ts, Us), NONE) + prfx [] narrowingsN (map narrowingT (Ts @ Us)) + |> Class.prove_instantiation_exit (K (Class.intro_classes_tac [])) + else + thy end; (* testing framework *) @@ -186,6 +199,7 @@ (** invocation of Haskell interpreter **) val narrowing_engine = File.read (Path.explode "~~/src/HOL/Tools/Quickcheck/Narrowing_Engine.hs") +val pnf_narrowing_engine = File.read (Path.explode "~~/src/HOL/Tools/Quickcheck/PNF_Narrowing_Engine.hs") fun exec verbose code = ML_Context.exec (fn () => Secure.use_text ML_Env.local_context (0, "generated code") verbose code) @@ -196,7 +210,7 @@ val _ = Isabelle_System.mkdirs path; in Exn.release (Exn.capture f path) end; -fun value ctxt (get, put, put_ml) (code, value_name) = +fun value contains_existentials ctxt (get, put, put_ml) (code, value_name) = let fun message s = if Config.get ctxt Quickcheck.quiet then () else Output.urgent_message s val tmp_prefix = "Quickcheck_Narrowing" @@ -216,16 +230,14 @@ val code' = prefix "module Code where {\n\ndata Typerep = Typerep String [Typerep];\n" (unprefix "module Code where {" code) val _ = File.write code_file code' - val _ = File.write narrowing_engine_file narrowing_engine + val _ = File.write narrowing_engine_file + (if contains_existentials then pnf_narrowing_engine else narrowing_engine) val _ = File.write main_file main val executable = File.shell_path (Path.append in_path (Path.basic "isabelle_quickcheck_narrowing")) val cmd = "exec \"$ISABELLE_GHC\" -fglasgow-exts " ^ (space_implode " " (map File.shell_path [code_file, narrowing_engine_file, main_file])) ^ " -o " ^ executable ^ ";" - val _ = if bash cmd <> 0 then - error "Compilation failed!" - else - () + val _ = if bash cmd <> 0 then error "Compilation with GHC failed" else () fun with_size k = if k > Config.get ctxt Quickcheck.size then NONE @@ -252,10 +264,10 @@ end in with_tmp_dir tmp_prefix run end; -fun dynamic_value_strict cookie thy postproc t = +fun dynamic_value_strict contains_existentials cookie thy postproc t = let val ctxt = Proof_Context.init_global thy - fun evaluator naming program ((_, vs_ty), t) deps = Exn.interruptible_capture (value ctxt cookie) + fun evaluator naming program ((_, vs_ty), t) deps = Exn.interruptible_capture (value contains_existentials ctxt cookie) (Code_Target.evaluator thy target naming program deps (vs_ty, t)); in Exn.release (Code_Thingol.dynamic_value thy (Exn.map_result o postproc) evaluator t) end; @@ -267,11 +279,29 @@ fun init _ () = error "Counterexample" ) +datatype counterexample = Universal_Counterexample of (term * counterexample) + | Existential_Counterexample of (term * counterexample) list + | Empty_Assignment + +fun map_counterexample f Empty_Assignment = Empty_Assignment + | map_counterexample f (Universal_Counterexample (t, c)) = + Universal_Counterexample (f t, map_counterexample f c) + | map_counterexample f (Existential_Counterexample cs) = + Existential_Counterexample (map (fn (t, c) => (f t, map_counterexample f c)) cs) + +structure Existential_Counterexample = Proof_Data +( + type T = unit -> counterexample option + fun init _ () = error "Counterexample" +) + +val put_existential_counterexample = Existential_Counterexample.put + val put_counterexample = Counterexample.put -fun finitize_functions t = +fun finitize_functions (xTs, t) = let - val ((names, Ts), t') = apfst split_list (strip_abs t) + val (names, boundTs) = split_list xTs fun mk_eval_ffun dT rT = Const (@{const_name "Quickcheck_Narrowing.eval_ffun"}, Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT]) --> dT --> rT) @@ -290,27 +320,97 @@ Type (@{type_name "Quickcheck_Narrowing.ffun"}, [dT, rT'])) end | eval_function T = (I, T) - val (tt, Ts') = split_list (map eval_function Ts) - val t'' = subst_bounds (map2 (fn f => fn x => f x) (rev tt) (map_index (Bound o fst) Ts), t') + val (tt, boundTs') = split_list (map eval_function boundTs) + val t' = subst_bounds (map2 (fn f => fn x => f x) (rev tt) (map_index (Bound o fst) boundTs), t) in - list_abs (names ~~ Ts', t'') + (names ~~ boundTs', t') end (** tester **) + +val rewrs = + map (swap o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) + (@{thms all_simps} @ @{thms ex_simps}) + @ map (HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) + [@{thm iff_conv_conj_imp}, @{thm not_ex}, @{thm not_all}] + +fun make_pnf_term thy t = Pattern.rewrite_term thy rewrs [] t + +fun strip_quantifiers (Const (@{const_name Ex}, _) $ Abs (x, T, t)) = + apfst (cons (@{const_name Ex}, (x, T))) (strip_quantifiers t) + | strip_quantifiers (Const (@{const_name All}, _) $ Abs (x, T, t)) = + apfst (cons (@{const_name All}, (x, T))) (strip_quantifiers t) + | strip_quantifiers t = ([], t) + +fun contains_existentials t = exists (fn (Q, _) => Q = @{const_name Ex}) (fst (strip_quantifiers t)) + +fun mk_property qs t = + let + fun enclose (@{const_name Ex}, (x, T)) t = + Const (@{const_name Quickcheck_Narrowing.exists}, (T --> @{typ property}) --> @{typ property}) + $ Abs (x, T, t) + | enclose (@{const_name All}, (x, T)) t = + Const (@{const_name Quickcheck_Narrowing.all}, (T --> @{typ property}) --> @{typ property}) + $ Abs (x, T, t) + in + fold_rev enclose qs (@{term Quickcheck_Narrowing.Property} $ + (list_comb (t , map Bound (((length qs) - 1) downto 0)))) + end + +fun mk_case_term ctxt p ((@{const_name Ex}, (x, T)) :: qs') (Existential_Counterexample cs) = + fst (Datatype.make_case ctxt Datatype_Case.Quiet [] (Free (x, T)) (map (fn (t, c) => + (t, mk_case_term ctxt (p - 1) qs' c)) cs)) + | mk_case_term ctxt p ((@{const_name All}, (x, T)) :: qs') (Universal_Counterexample (t, c)) = + if p = 0 then t else mk_case_term ctxt (p - 1) qs' c + +fun mk_terms ctxt qs result = + let + val + ps = filter (fn (_, (@{const_name All}, _)) => true | _ => false) (map_index I qs) + in + map (fn (p, (_, (x, T))) => (x, mk_case_term ctxt p qs result)) ps + end -fun test_term ctxt (limit_time, is_interactive) (t, eval_terms) = +fun test_term ctxt (limit_time, is_interactive) (t, eval_terms) = let val thy = Proof_Context.theory_of ctxt - val t' = list_abs_free (Term.add_frees t [], t) - val t'' = if Config.get ctxt finite_functions then finitize_functions t' else t' - fun ensure_testable t = - Const (@{const_name Quickcheck_Narrowing.ensure_testable}, fastype_of t --> fastype_of t) $ t - val result = dynamic_value_strict - (Counterexample.get, Counterexample.put, "Narrowing_Generators.put_counterexample") - thy (Option.map o map) (ensure_testable t'') + val t' = fold_rev (fn (x, T) => fn t => HOLogic.mk_all (x, T, t)) (Term.add_frees t []) t + val pnf_t = make_pnf_term thy t' in - Quickcheck.Result {counterexample = Option.map ((curry (op ~~)) (Term.add_free_names t [])) result, - evaluation_terms = Option.map (K []) result, timings = [], reports = []} + if Config.get ctxt allow_existentials andalso contains_existentials pnf_t then + let + fun wrap f (qs, t) = + let val (qs1, qs2) = split_list qs in + apfst (map2 pair qs1) (f (qs2, t)) end + val finitize = if Config.get ctxt finite_functions then wrap finitize_functions else I + val (qs, prop_t) = finitize (strip_quantifiers pnf_t) + val prop_term = fold_rev (fn (_, (x, T)) => fn t => Abs (x, T, t)) qs prop_t + val ((prop_def, _), ctxt') = Local_Theory.define ((Binding.conceal (Binding.name "test_property"), NoSyn), + ((Binding.conceal Binding.empty, [Code.add_default_eqn_attrib]), prop_term)) ctxt + val (prop_def', thy') = Local_Theory.exit_result_global Morphism.term (prop_def, ctxt') + val result = dynamic_value_strict true + (Existential_Counterexample.get, Existential_Counterexample.put, + "Narrowing_Generators.put_existential_counterexample") + thy' (Option.map o map_counterexample) (mk_property qs prop_def') + val result' = Option.map (mk_terms ctxt' (fst (strip_quantifiers pnf_t))) result + in + Quickcheck.Result {counterexample = result', evaluation_terms = Option.map (K []) result, + timings = [], reports = []} + end + else + let + val t' = Term.list_abs_free (Term.add_frees t [], t) + fun wrap f t = list_abs (f (strip_abs t)) + val finitize = if Config.get ctxt finite_functions then wrap finitize_functions else I + fun ensure_testable t = + Const (@{const_name Quickcheck_Narrowing.ensure_testable}, fastype_of t --> fastype_of t) $ t + val result = dynamic_value_strict false + (Counterexample.get, Counterexample.put, "Narrowing_Generators.put_counterexample") + thy (Option.map o map) (ensure_testable (finitize t')) + in + Quickcheck.Result {counterexample = Option.map ((curry (op ~~)) (Term.add_free_names t [])) result, + evaluation_terms = Option.map (K []) result, timings = [], reports = []} + end end; fun test_goals ctxt (limit_time, is_interactive) insts goals = diff -r 9d717505595f -r 3c58977e0911 src/HOL/ex/Birthday_Paradox.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Birthday_Paradox.thy Tue Jun 07 11:24:16 2011 +0200 @@ -0,0 +1,101 @@ +(* Title: HOL/ex/Birthday_Paradox.thy + Author: Lukas Bulwahn, TU Muenchen, 2007 +*) + +header {* A Formulation of the Birthday Paradox *} + +theory Birthday_Paradox +imports Main "~~/src/HOL/Fact" "~~/src/HOL/Library/FuncSet" +begin + +section {* Cardinality *} + +lemma card_product_dependent: + assumes "finite S" + assumes "\x \ S. finite (T x)" + shows "card {(x, y). x \ S \ y \ T x} = (\x \ S. card (T x))" +proof - + note `finite S` + moreover + have "{(x, y). x \ S \ y \ T x} = (UN x : S. Pair x ` T x)" by auto + moreover + from `\x \ S. finite (T x)` have "ALL x:S. finite (Pair x ` T x)" by auto + moreover + have " ALL i:S. ALL j:S. i ~= j --> Pair i ` T i Int Pair j ` T j = {}" by auto + moreover + ultimately have "card {(x, y). x \ S \ y \ T x} = (SUM i:S. card (Pair i ` T i))" + by (auto, subst card_UN_disjoint) auto + also have "... = (SUM x:S. card (T x))" + by (subst card_image) (auto intro: inj_onI) + finally show ?thesis by auto +qed + +lemma card_extensional_funcset_inj_on: + assumes "finite S" "finite T" "card S \ card T" + shows "card {f \ extensional_funcset S T. inj_on f S} = fact (card T) div (fact (card T - card S))" +using assms +proof (induct S arbitrary: T rule: finite_induct) + case empty + from this show ?case by (simp add: Collect_conv_if extensional_funcset_empty_domain) +next + case (insert x S) + { fix x + from `finite T` have "finite (T - {x})" by auto + from `finite S` this have "finite (extensional_funcset S (T - {x}))" + by (rule finite_extensional_funcset) + moreover + have "{f : extensional_funcset S (T - {x}). inj_on f S} \ (extensional_funcset S (T - {x}))" by auto + ultimately have "finite {f : extensional_funcset S (T - {x}). inj_on f S}" + by (auto intro: finite_subset) + } note finite_delete = this + from insert have hyps: "\y \ T. card ({g. g \ extensional_funcset S (T - {y}) \ inj_on g S}) = fact (card T - 1) div fact ((card T - 1) - card S)"(is "\ _ \ T. _ = ?k") by auto + from extensional_funcset_extend_domain_inj_on_eq[OF `x \ S`] + have "card {f. f : extensional_funcset (insert x S) T & inj_on f (insert x S)} = + card ((%(y, g). g(x := y)) ` {(y, g). y : T & g : extensional_funcset S (T - {y}) & inj_on g S})" + by metis + also from extensional_funcset_extend_domain_inj_onI[OF `x \ S`, of T] have "... = card {(y, g). y : T & g : extensional_funcset S (T - {y}) & inj_on g S}" + by (simp add: card_image) + also have "card {(y, g). y \ T \ g \ extensional_funcset S (T - {y}) \ inj_on g S} = + card {(y, g). y \ T \ g \ {f \ extensional_funcset S (T - {y}). inj_on f S}}" by auto + also from `finite T` finite_delete have "... = (\y \ T. card {g. g \ extensional_funcset S (T - {y}) \ inj_on g S})" + by (subst card_product_dependent) auto + also from hyps have "... = (card T) * ?k" + by auto + also have "... = card T * fact (card T - 1) div fact (card T - card (insert x S))" + using insert unfolding div_mult1_eq[of "card T" "fact (card T - 1)"] + by (simp add: fact_mod) + also have "... = fact (card T) div fact (card T - card (insert x S))" + using insert by (simp add: fact_reduce_nat[of "card T"]) + finally show ?case . +qed + +lemma card_extensional_funcset_not_inj_on: + assumes "finite S" "finite T" "card S \ card T" + shows "card {f \ extensional_funcset S T. \ inj_on f S} = (card T) ^ (card S) - (fact (card T)) div (fact (card T - card S))" +proof - + have subset: "{f : extensional_funcset S T. inj_on f S} <= extensional_funcset S T" by auto + from finite_subset[OF subset] assms have finite: "finite {f : extensional_funcset S T. inj_on f S}" + by (auto intro!: finite_extensional_funcset) + have "{f \ extensional_funcset S T. \ inj_on f S} = extensional_funcset S T - {f \ extensional_funcset S T. inj_on f S}" by auto + from assms this finite subset show ?thesis + by (simp add: card_Diff_subset card_extensional_funcset card_extensional_funcset_inj_on) +qed + +lemma setprod_upto_nat_unfold: + "setprod f {m..(n::nat)} = (if n < m then 1 else (if n = 0 then f 0 else f n * setprod f {m..(n - 1)}))" + by auto (auto simp add: gr0_conv_Suc atLeastAtMostSuc_conv) + +section {* Birthday paradox *} + +lemma birthday_paradox: + assumes "card S = 23" "card T = 365" + shows "2 * card {f \ extensional_funcset S T. \ inj_on f S} \ card (extensional_funcset S T)" +proof - + from `card S = 23` `card T = 365` have "finite S" "finite T" "card S <= card T" by (auto intro: card_ge_0_finite) + from assms show ?thesis + using card_extensional_funcset[OF `finite S`, of T] + card_extensional_funcset_not_inj_on[OF `finite S` `finite T` `card S <= card T`] + by (simp add: fact_div_fact setprod_upto_nat_unfold) +qed + +end diff -r 9d717505595f -r 3c58977e0911 src/HOL/ex/Birthday_Paradoxon.thy --- a/src/HOL/ex/Birthday_Paradoxon.thy Tue Jun 07 11:05:09 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,101 +0,0 @@ -(* Title: HOL/ex/Birthday_Paradoxon.thy - Author: Lukas Bulwahn, TU Muenchen, 2007 -*) - -header {* A Formulation of the Birthday Paradoxon *} - -theory Birthday_Paradoxon -imports Main "~~/src/HOL/Fact" "~~/src/HOL/Library/FuncSet" -begin - -section {* Cardinality *} - -lemma card_product_dependent: - assumes "finite S" - assumes "\x \ S. finite (T x)" - shows "card {(x, y). x \ S \ y \ T x} = (\x \ S. card (T x))" -proof - - note `finite S` - moreover - have "{(x, y). x \ S \ y \ T x} = (UN x : S. Pair x ` T x)" by auto - moreover - from `\x \ S. finite (T x)` have "ALL x:S. finite (Pair x ` T x)" by auto - moreover - have " ALL i:S. ALL j:S. i ~= j --> Pair i ` T i Int Pair j ` T j = {}" by auto - moreover - ultimately have "card {(x, y). x \ S \ y \ T x} = (SUM i:S. card (Pair i ` T i))" - by (auto, subst card_UN_disjoint) auto - also have "... = (SUM x:S. card (T x))" - by (subst card_image) (auto intro: inj_onI) - finally show ?thesis by auto -qed - -lemma card_extensional_funcset_inj_on: - assumes "finite S" "finite T" "card S \ card T" - shows "card {f \ extensional_funcset S T. inj_on f S} = fact (card T) div (fact (card T - card S))" -using assms -proof (induct S arbitrary: T rule: finite_induct) - case empty - from this show ?case by (simp add: Collect_conv_if extensional_funcset_empty_domain) -next - case (insert x S) - { fix x - from `finite T` have "finite (T - {x})" by auto - from `finite S` this have "finite (extensional_funcset S (T - {x}))" - by (rule finite_extensional_funcset) - moreover - have "{f : extensional_funcset S (T - {x}). inj_on f S} \ (extensional_funcset S (T - {x}))" by auto - ultimately have "finite {f : extensional_funcset S (T - {x}). inj_on f S}" - by (auto intro: finite_subset) - } note finite_delete = this - from insert have hyps: "\y \ T. card ({g. g \ extensional_funcset S (T - {y}) \ inj_on g S}) = fact (card T - 1) div fact ((card T - 1) - card S)"(is "\ _ \ T. _ = ?k") by auto - from extensional_funcset_extend_domain_inj_on_eq[OF `x \ S`] - have "card {f. f : extensional_funcset (insert x S) T & inj_on f (insert x S)} = - card ((%(y, g). g(x := y)) ` {(y, g). y : T & g : extensional_funcset S (T - {y}) & inj_on g S})" - by metis - also from extensional_funcset_extend_domain_inj_onI[OF `x \ S`, of T] have "... = card {(y, g). y : T & g : extensional_funcset S (T - {y}) & inj_on g S}" - by (simp add: card_image) - also have "card {(y, g). y \ T \ g \ extensional_funcset S (T - {y}) \ inj_on g S} = - card {(y, g). y \ T \ g \ {f \ extensional_funcset S (T - {y}). inj_on f S}}" by auto - also from `finite T` finite_delete have "... = (\y \ T. card {g. g \ extensional_funcset S (T - {y}) \ inj_on g S})" - by (subst card_product_dependent) auto - also from hyps have "... = (card T) * ?k" - by auto - also have "... = card T * fact (card T - 1) div fact (card T - card (insert x S))" - using insert unfolding div_mult1_eq[of "card T" "fact (card T - 1)"] - by (simp add: fact_mod) - also have "... = fact (card T) div fact (card T - card (insert x S))" - using insert by (simp add: fact_reduce_nat[of "card T"]) - finally show ?case . -qed - -lemma card_extensional_funcset_not_inj_on: - assumes "finite S" "finite T" "card S \ card T" - shows "card {f \ extensional_funcset S T. \ inj_on f S} = (card T) ^ (card S) - (fact (card T)) div (fact (card T - card S))" -proof - - have subset: "{f : extensional_funcset S T. inj_on f S} <= extensional_funcset S T" by auto - from finite_subset[OF subset] assms have finite: "finite {f : extensional_funcset S T. inj_on f S}" - by (auto intro!: finite_extensional_funcset) - have "{f \ extensional_funcset S T. \ inj_on f S} = extensional_funcset S T - {f \ extensional_funcset S T. inj_on f S}" by auto - from assms this finite subset show ?thesis - by (simp add: card_Diff_subset card_extensional_funcset card_extensional_funcset_inj_on) -qed - -lemma setprod_upto_nat_unfold: - "setprod f {m..(n::nat)} = (if n < m then 1 else (if n = 0 then f 0 else f n * setprod f {m..(n - 1)}))" - by auto (auto simp add: gr0_conv_Suc atLeastAtMostSuc_conv) - -section {* Birthday paradoxon *} - -lemma birthday_paradoxon: - assumes "card S = 23" "card T = 365" - shows "2 * card {f \ extensional_funcset S T. \ inj_on f S} \ card (extensional_funcset S T)" -proof - - from `card S = 23` `card T = 365` have "finite S" "finite T" "card S <= card T" by (auto intro: card_ge_0_finite) - from assms show ?thesis - using card_extensional_funcset[OF `finite S`, of T] - card_extensional_funcset_not_inj_on[OF `finite S` `finite T` `card S <= card T`] - by (simp add: fact_div_fact setprod_upto_nat_unfold) -qed - -end diff -r 9d717505595f -r 3c58977e0911 src/HOL/ex/Quickcheck_Narrowing_Examples.thy --- a/src/HOL/ex/Quickcheck_Narrowing_Examples.thy Tue Jun 07 11:05:09 2011 +0200 +++ b/src/HOL/ex/Quickcheck_Narrowing_Examples.thy Tue Jun 07 11:24:16 2011 +0200 @@ -23,14 +23,56 @@ oops *) +(*declare [[quickcheck_narrowing_overlord]]*) +subsection {* Simple examples with existentials *} + +lemma + "\ y :: nat. \ x. x = y" +quickcheck[tester = narrowing, finite_types = false, expect = counterexample] +oops + +lemma + "x > 1 --> (\y :: nat. x < y & y <= 1)" +quickcheck[tester = narrowing, finite_types = false, expect = counterexample] +oops + +lemma + "x > 2 --> (\y :: nat. x < y & y <= 2)" +quickcheck[tester = narrowing, finite_types = false, size = 10] +oops + +lemma + "\ x. \ y :: nat. x > 3 --> (y < x & y > 3)" +quickcheck[tester = narrowing, finite_types = false, size = 7] +oops + +text {* A false conjecture derived from an partial copy-n-paste of @{thm not_distinct_decomp} *} +lemma + "~ distinct ws ==> EX xs ys zs y. ws = xs @ [y] @ ys @ [y]" +quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] +oops + +text {* A false conjecture derived from theorems @{thm split_list_first} and @{thm split_list_last} *} +lemma + "x : set xs ==> EX ys zs. xs = ys @ x # zs & x ~: set zs & x ~: set ys" +quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] +oops + +text {* A false conjecture derived from @{thm map_eq_Cons_conv} with a typo *} +lemma + "(map f xs = y # ys) = (EX z zs. xs = z' # zs & f z = y & map f zs = ys)" +quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] +oops + subsection {* Simple list examples *} lemma "rev xs = xs" - quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] +quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] oops +(* FIXME: integer has strange representation! *) lemma "rev xs = xs" - quickcheck[tester = narrowing, finite_types = false, default_type = int, expect = counterexample] +quickcheck[tester = narrowing, finite_types = false, default_type = int, expect = counterexample] oops (* lemma "rev xs = xs" @@ -39,30 +81,29 @@ *) subsection {* Simple examples with functions *} -declare [[quickcheck_finite_functions]] -(* lemma "map f xs = map g xs" - quickcheck[tester = narrowing, finite_types = true, expect = counterexample] + quickcheck[tester = narrowing, finite_types = false, expect = counterexample] oops lemma "map f xs = map f ys ==> xs = ys" - quickcheck[tester = narrowing, finite_types = true, expect = counterexample] + quickcheck[tester = narrowing, finite_types = false, expect = counterexample] oops lemma "list_all2 P (rev xs) (rev ys) = list_all2 P xs (rev ys)" - quickcheck[tester = narrowing, expect = counterexample] + quickcheck[tester = narrowing, finite_types = false, expect = counterexample] oops lemma "map f xs = F f xs" - quickcheck[tester = narrowing, finite_types = true, expect = counterexample] -oops -*) -lemma "map f xs = F f xs" quickcheck[tester = narrowing, finite_types = false, expect = counterexample] oops -(* requires some work... +lemma "map f xs = F f xs" + quickcheck[tester = narrowing, finite_types = false, expect = counterexample] +oops + +(* requires some work...*) +(* lemma "R O S = S O R" quickcheck[tester = narrowing, size = 2] oops diff -r 9d717505595f -r 3c58977e0911 src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Tue Jun 07 11:05:09 2011 +0200 +++ b/src/HOL/ex/ROOT.ML Tue Jun 07 11:24:16 2011 +0200 @@ -73,7 +73,7 @@ "Gauge_Integration", "Dedekind_Real", "Quicksort", - "Birthday_Paradoxon", + "Birthday_Paradox", "List_to_Set_Comprehension_Examples", "Set_Algebras" ];