# HG changeset patch # User paulson # Date 1596537903 -3600 # Node ID 3f8e6c0166ac0d644640f76f6b9e3fa3b55e8be8 # Parent 2030eacf3a723dda2d94b6da8e43fcf4bda1ff7d# Parent b6065cbbf5e29aa281122f2f40d0c980c9cf1641 merged diff -r 2030eacf3a72 -r 3f8e6c0166ac src/HOL/Groups_Big.thy --- a/src/HOL/Groups_Big.thy Mon Aug 03 16:21:33 2020 +0200 +++ b/src/HOL/Groups_Big.thy Tue Aug 04 11:45:03 2020 +0100 @@ -165,6 +165,11 @@ shows "F g A = F h B" using assms by (simp add: reindex) +lemma image_eq: + assumes "inj_on g A" + shows "F (\x. x) (g ` A) = F g A" + using assms reindex_cong by fastforce + lemma UNION_disjoint: assumes "finite I" and "\i\I. finite (A i)" and "\i\I. \j\I. i \ j \ A i \ A j = {}" diff -r 2030eacf3a72 -r 3f8e6c0166ac src/HOL/Library/Infinite_Set.thy --- a/src/HOL/Library/Infinite_Set.thy Mon Aug 03 16:21:33 2020 +0200 +++ b/src/HOL/Library/Infinite_Set.thy Tue Aug 04 11:45:03 2020 +0100 @@ -484,7 +484,7 @@ lemma finite_enumerate_initial_segment: fixes S :: "'a::wellorder set" - assumes "finite S" "s \ S" and n: "n < card (S \ {.. {.. {.. {s \ S. enumerate S n < s}" + obtain T where T: "T \ {s \ S. enumerate S n < s}" by (metis Infinite_Set.enumerate_step enumerate_in_set finite_enumerate_in_set finite_enumerate_step less_card mem_Collect_eq) - have "(LEAST x. x \ S \ x < s \ enumerate S n < x) = (LEAST s. s \ S \ enumerate S n < s)" + have "(LEAST x. x \ S \ x < s \ enumerate S n < x) = (LEAST x. x \ S \ enumerate S n < x)" (is "_ = ?r") proof (intro Least_equality conjI) show "?r \ S" - by (metis (mono_tags, lifting) LeastI mem_Collect_eq t) - show "?r < s" - using not_less_Least [of _ "\t. t \ S \ enumerate S n < t"] Suc assms - by (metis (no_types, lifting) Int_Collect Suc_lessD finite_Int finite_enumerate_in_set finite_enumerate_step lessThan_def linorder_cases) + by (metis (mono_tags, lifting) LeastI mem_Collect_eq T) + have "\ s \ ?r" + using not_less_Least [of _ "\x. x \ S \ enumerate S n < x"] Suc assms + by (metis (mono_tags, lifting) Int_Collect Suc_lessD finite_Int finite_enumerate_in_set finite_enumerate_step lessThan_def less_le_trans) + then show "?r < s" + by auto show "enumerate S n < ?r" - by (metis (no_types, lifting) LeastI mem_Collect_eq t) + by (metis (no_types, lifting) LeastI mem_Collect_eq T) qed (auto simp: Least_le) then show ?case using Suc assms by (simp add: finite_enumerate_Suc'' less_card)