merged
authorwenzelm
Thu, 30 Jul 2015 21:56:19 +0200
changeset 60836 c5db501da8e4
parent 60828 e9e272fa8daf (diff)
parent 60835 6512bb0b1ff4 (current diff)
child 60837 c362049f3f84
child 60838 2d7eea27ceec
merged
--- a/src/HOL/Library/Infinite_Set.thy	Thu Jul 30 14:02:19 2015 +0200
+++ b/src/HOL/Library/Infinite_Set.thy	Thu Jul 30 21:56:19 2015 +0200
@@ -62,6 +62,16 @@
   with S show False by simp
 qed
 
+lemma infinite_coinduct [consumes 1, case_names infinite]:
+  assumes "X A"
+  and step: "\<And>A. X A \<Longrightarrow> \<exists>x\<in>A. X (A - {x}) \<or> infinite (A - {x})"
+  shows "infinite A"
+proof
+  assume "finite A"
+  then show False using \<open>X A\<close>
+    by(induction rule: finite_psubset_induct)(meson Diff_subset card_Diff1_less card_psubset finite_Diff step)
+qed    
+
 text \<open>
   As a concrete example, we prove that the set of natural numbers is
   infinite.