# HG changeset patch # User nipkow # Date 1667481607 -3600 # Node ID 2612b3406b61242f34d225e012f0d2cc071996d6 # Parent 2fd70eb1e9b61df991e3349ea972a4d40483b48b added finite simproc diff -r 2fd70eb1e9b6 -r 2612b3406b61 src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Wed Nov 02 18:58:38 2022 +0100 +++ b/src/HOL/Finite_Set.thy Thu Nov 03 14:20:07 2022 +0100 @@ -193,6 +193,35 @@ lemma finite_subset: "A \ B \ finite B \ finite A" by (rule rev_finite_subset) +simproc_setup finite ("finite A") = \fn _ => +let + val finite_subset = @{thm finite_subset} + val Eq_TrueI = @{thm Eq_TrueI} + + fun is_subset A th = case Thm.prop_of th of + (_ $ (Const (\<^const_name>\less_eq\, Type (\<^type_name>\fun\, [Type (\<^type_name>\set\, _), _])) $ A' $ B)) + => if A aconv A' then SOME(B,th) else NONE + | _ => NONE; + + fun is_finite th = case Thm.prop_of th of + (_ $ (Const (\<^const_name>\finite\, _) $ A)) => SOME(A,th) + | _ => NONE; + + fun comb (A,sub_th) (A',fin_th) ths = if A aconv A' then (sub_th,fin_th) :: ths else ths + + fun proc ss ct = + (let + val _ $ A = Thm.term_of ct + val prems = Simplifier.prems_of ss + val fins = map_filter is_finite prems + val subsets = map_filter (is_subset A) prems + in case fold_product comb subsets fins [] of + (sub_th,fin_th) :: _ => SOME((fin_th RS (sub_th RS finite_subset)) RS Eq_TrueI) + | _ => NONE + end) +in proc end +\ + lemma finite_UnI: assumes "finite F" and "finite G" shows "finite (F \ G)"