# HG changeset patch # User paulson # Date 1137170381 -3600 # Node ID 01265301db7f013a38a5e7cb64e4f9f0e6495dcf # Parent 5bce9fddce2ea349fa65785af61afa03aa70dd78 blacklist experiments diff -r 5bce9fddce2e -r 01265301db7f src/HOL/Tools/ATP/res_clasimpset.ML --- a/src/HOL/Tools/ATP/res_clasimpset.ML Fri Jan 13 17:39:19 2006 +0100 +++ b/src/HOL/Tools/ATP/res_clasimpset.ML Fri Jan 13 17:39:41 2006 +0100 @@ -75,6 +75,7 @@ signature RES_CLASIMP = sig val blacklist : string list ref (*Theorems forbidden in the output*) + val theory_blacklist : string list ref (*entire blacklisted theories*) val relevant : int ref val use_simpset: bool ref val get_clasimp_lemmas : @@ -87,55 +88,173 @@ val use_simpset = ref false; (*Performance is much better without simprules*) (*In general, these produce clauses that are prolific (match too many equality or - membership literals) and relate to seldom-used facts. Some duplicate other rules.*) + membership literals) and relate to seldom-used facts. Some duplicate other rules. + FIXME: this blacklist needs to be maintained using theory data and added to using + an attribute.*) val blacklist = ref - ["Finite_Set.Max_ge", + ["Datatype.not_None_eq_iff2", + "Datatype.not_Some_eq_D", + "Datatype.not_Some_eq", + "Datatype.option.size_1", + "Datatype.option.size_2", + "Datatype.prod.size", + "Datatype.sum.size_1", + "Datatype.sum.size_2", + "Datatype.unit.size", + "Divides.dvd_0_left_iff", + "Finite_Set.card_0_eq", + "Finite_Set.card_infinite", + "Finite_Set.Max_ge", + "Finite_Set.Max_in", "Finite_Set.Max_le_iff", "Finite_Set.Max_less_iff", - "Finite_Set.Min_le", + "Finite_Set.max.f_below_strict_below.below_f_conv", (*duplicates in Orderings.*) + "Finite_Set.max.f_below_strict_below.strict_below_f_conv", (*duplicates in Orderings.*) "Finite_Set.Min_ge_iff", "Finite_Set.Min_gr_iff", - "Finite_Set.min_max.below_inf_sup_Inf_Sup.inf_Sup_absorb", - "Finite_Set.min_max.below_inf_sup_Inf_Sup.sup_Inf_absorb", - (*The next four are duplicates of the properties of min and max, from Orderings.*) - "Finite_Set.max.f_below_strict_below.below_f_conv", - "Finite_Set.max.f_below_strict_below.strict_below_f_conv", - "Finite_Set.min.f_below_strict_below.strict_below_f_conv", - "Finite_Set.min.f_below_strict_below.below_f_conv", + "Finite_Set.Min_in", + "Finite_Set.Min_le", + "Finite_Set.min_max.below_inf_sup_Inf_Sup.inf_Sup_absorb", + "Finite_Set.min_max.below_inf_sup_Inf_Sup.sup_Inf_absorb", + "Finite_Set.min.f_below_strict_below.below_f_conv", (*duplicates in Orderings.*) + "Finite_Set.min.f_below_strict_below.strict_below_f_conv", (*duplicates in Orderings.*) "Infinite_Set.atmost_one_unique", + "IntArith.zabs_less_one_iff", + "IntDef.Integ.Abs_Integ_inject", + "IntDef.Integ.Abs_Integ_inverse", + "IntDiv.zdvd_0_left", + "IntDiv.zero_less_zpower_abs_iff", + "List.append_eq_append_conv", + "List.Cons_in_lex", "List.in_listsD", "List.in_listsI", + "List.lists.Cons", "List.listsE", - "List.lists.Cons", + "List.take_eq_Nil", + "Nat.less_one", + "Nat.less_one", (*not directional? obscure*) + "Nat.not_gr0", + "Nat.one_eq_mult_iff", (*duplicate by symmetry*) + "NatArith.of_nat_0_eq_iff", + "NatArith.of_nat_eq_0_iff", + "NatArith.of_nat_le_0_iff", + "NatSimprocs.divide_le_0_iff_number_of", (*seldom used; often prolific*) + "NatSimprocs.divide_le_0_iff_number_of", (*too many clauses*) + "NatSimprocs.divide_less_0_iff_number_of", + "NatSimprocs.divide_less_0_iff_number_of", (*too many clauses*) + "NatSimprocs.equation_minus_iff_1", (*not directional*) + "NatSimprocs.equation_minus_iff_number_of", (*not directional*) + "NatSimprocs.le_minus_iff_1", (*not directional*) + "NatSimprocs.le_minus_iff_number_of", (*not directional*) + "NatSimprocs.less_minus_iff_1", (*not directional*) + "NatSimprocs.less_minus_iff_number_of", (*not directional*) + "NatSimprocs.minus_equation_iff_number_of", (*not directional*) + "NatSimprocs.minus_le_iff_1", (*not directional*) + "NatSimprocs.minus_le_iff_number_of", (*not directional*) + "NatSimprocs.minus_less_iff_1", (*not directional*) + "NatSimprocs.mult_le_cancel_left_number_of", (*excessive case analysis*) + "NatSimprocs.mult_le_cancel_right_number_of", (*excessive case analysis*) + "NatSimprocs.mult_less_cancel_left_number_of", (*excessive case analysis*) + "NatSimprocs.mult_less_cancel_right_number_of", (*excessive case analysis*) + "NatSimprocs.zero_le_divide_iff_number_of", + "NatSimprocs.zero_le_divide_iff_number_of", (*excessive case analysis*) + "NatSimprocs.zero_less_divide_iff_number_of", + "NatSimprocs.zero_less_divide_iff_number_of", (*excessive case analysis*) + "OrderedGroup.abs_0_eq", + "OrderedGroup.abs_0_eq", (*duplicate by symmetry*) + "OrderedGroup.diff_eq_0_iff_eq", (*prolific?*) + "OrderedGroup.join_0_eq_0", + "OrderedGroup.meet_0_eq_0", + "OrderedGroup.pprt_eq_0", (*obscure*) + "OrderedGroup.pprt_eq_id", (*obscure*) + "OrderedGroup.pprt_mono", (*obscure*) + "Parity.even_nat_power", (*obscure, somewhat prolilfic*) + "Parity.power_eq_0_iff_number_of", + "Parity.power_eq_0_iff_number_of", + "Parity.power_le_zero_eq_number_of", + "Parity.power_le_zero_eq_number_of", (*obscure and prolific*) + "Parity.power_less_zero_eq_number_of", + "Parity.zero_le_power_eq_number_of", (*obscure and prolific*) + "Parity.zero_less_power_eq_number_of", (*obscure and prolific*) + "Power.zero_less_power_abs_iff", + "Relation.diagI", "Relation.ImageI", - "Relation.diagI", + "Ring_and_Field.divide_cancel_left", (*fields are seldom used & often prolific*) + "Ring_and_Field.divide_cancel_right", + "Ring_and_Field.divide_divide_eq_left", + "Ring_and_Field.divide_divide_eq_right", + "Ring_and_Field.divide_eq_0_iff", + "Ring_and_Field.divide_eq_1_iff", + "Ring_and_Field.divide_eq_eq_1", + "Ring_and_Field.divide_le_0_1_iff", + "Ring_and_Field.divide_le_eq_1_neg", + "Ring_and_Field.divide_le_eq_1_neg", (*obscure and prolific*) + "Ring_and_Field.divide_le_eq_1_pos", + "Ring_and_Field.divide_le_eq_1_pos", (*obscure and prolific*) + "Ring_and_Field.divide_less_0_1_iff", + "Ring_and_Field.divide_less_eq_1_neg", (*obscure and prolific*) + "Ring_and_Field.divide_less_eq_1_pos", + "Ring_and_Field.divide_less_eq_1_pos", (*obscure and prolific*) + "Ring_and_Field.eq_divide_eq_1", + "Ring_and_Field.eq_divide_eq_1", (*duplicate by symmetry*) + "Ring_and_Field.field_mult_cancel_left", + "Ring_and_Field.field_mult_cancel_right", + "Ring_and_Field.inverse_le_iff_le_neg", + "Ring_and_Field.inverse_le_iff_le", + "Ring_and_Field.inverse_less_iff_less_neg", + "Ring_and_Field.inverse_less_iff_less", + "Ring_and_Field.le_divide_eq_1_neg", + "Ring_and_Field.le_divide_eq_1_neg", (*obscure and prolific*) + "Ring_and_Field.le_divide_eq_1_pos", + "Ring_and_Field.le_divide_eq_1_pos", (*obscure and prolific*) + "Ring_and_Field.less_divide_eq_1_neg", + "Ring_and_Field.less_divide_eq_1_neg", (*obscure and prolific*) + "Ring_and_Field.less_divide_eq_1_pos", + "Ring_and_Field.less_divide_eq_1_pos", (*obscure and prolific*) + "Ring_and_Field.one_eq_divide_iff", (*duplicate by symmetry*) + "Set.Diff_eq_empty_iff", (*redundant with paramodulation*) "Set.Diff_insert0", + "Set.disjoint_insert_1", + "Set.disjoint_insert_2", + "Set.empty_Union_conv", (*redundant with paramodulation*) + "Set.insert_disjoint_1", + "Set.insert_disjoint_2", + "Set.Int_UNIV", (*redundant with paramodulation*) + "Set.Inter_iff", (*We already have InterI, InterE*) "Set.Inter_UNIV_conv_1", "Set.Inter_UNIV_conv_2", - "Set.Inter_iff", (*We already have InterI, InterE*) + "Set.psubsetE", (*too prolific and obscure*) + "Set.psubsetI", + "Set.singleton_insert_inj_eq'", + "Set.singleton_insert_inj_eq", + "Set.singletonD", (*these two duplicate some "insert" lemmas*) + "Set.singletonI", + "Set.Un_empty", (*redundant with paramodulation*) + "Set.Union_empty_conv", (*redundant with paramodulation*) "Set.Union_iff", (*We already have UnionI, UnionE*) - "Datatype.not_None_eq_iff2", - "Datatype.not_Some_eq", - "Datatype.not_Some_eq_D", - "Set.disjoint_insert_1", - "Set.disjoint_insert_2", - "Set.insert_disjoint_1", - "Set.insert_disjoint_2", - "Set.singletonD", - "Set.singletonI", + "SetInterval.atLeastAtMost_iff", (*obscure and prolific*) + "SetInterval.atLeastLessThan_iff", (*obscure and prolific*) + "SetInterval.greaterThanAtMost_iff", (*obscure and prolific*) + "SetInterval.greaterThanLessThan_iff", (*obscure and prolific*) + "SetInterval.ivl_subset", (*excessive case analysis*) "Sum_Type.InlI", - "Sum_Type.InrI", - "Set.singleton_insert_inj_eq", - "Set.singleton_insert_inj_eq'"]; - -(*These are the defining properties of min and max, but as they are prolific they - could be included above. - "Orderings.max_less_iff_conj", + "Sum_Type.InrI"]; + +(*These might be prolific but are probably OK, and min and max are basic. + "Orderings.max_less_iff_conj", "Orderings.min_less_iff_conj", "Orderings.min_max.below_inf.below_inf_conv", "Orderings.min_max.below_sup.above_sup_conv", +Very prolific and somewhat obscure: + "Set.InterD", + "Set.UnionI", *) +val theory_blacklist = ref + ["Datatype_Universe.", (*unnecessary in virtually all proofs*) + "Equiv_Relations."] + + val relevant = ref 0; (*Relevance filtering is off by default*) (*The "name" of a theorem is its statement, if nothing else is available.*) @@ -165,9 +284,7 @@ multi (clause, theorem) num_of_cls [] end; -(*Hashing to detect duplicate and variant clauses, e.g. from the [iff] attribute -Some primes from http://primes.utm.edu/: -*) +(*Hashing to detect duplicate and variant clauses, e.g. from the [iff] attribute*) exception HASH_CLAUSE and HASH_STRING; @@ -177,10 +294,12 @@ Polyhash.insert ht (x^"_iff2", ()); Polyhash.insert ht (x^"_dest", ())); +fun banned_theory s = exists (fn thy => String.isPrefix thy s) (!theory_blacklist); + fun make_banned_test xs = let val ht = Polyhash.mkTable (Polyhash.hash_string, op =) (6000, HASH_STRING) - fun banned s = isSome (Polyhash.peek ht s) + fun banned s = isSome (Polyhash.peek ht s) orelse banned_theory s in app (fn x => Polyhash.insert ht (x,())) (!blacklist); app (insert_suffixed_names ht) (!blacklist @ xs); banned