merged
authorhuffman
Thu, 18 Aug 2011 14:08:39 -0700
changeset 44283 637297cf6142
parent 44281 d27b9fe4759e (diff)
parent 44282 f0de18b62d63 (current diff)
child 44284 584a590ce6d3
merged
--- a/src/HOL/Equiv_Relations.thy	Thu Aug 18 13:36:58 2011 -0700
+++ b/src/HOL/Equiv_Relations.thy	Thu Aug 18 14:08:39 2011 -0700
@@ -164,7 +164,7 @@
 
 text{*A congruence-preserving function*}
 
-definition congruent :: "('a \<times> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"  where
+definition congruent :: "('a \<times> 'a) set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"  where
   "congruent r f \<longleftrightarrow> (\<forall>(y, z) \<in> r. f y = f z)"
 
 lemma congruentI:
@@ -229,7 +229,7 @@
 
 text{*A congruence-preserving function of two arguments*}
 
-definition congruent2 :: "('a \<times> 'a \<Rightarrow> bool) \<Rightarrow> ('b \<times> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> bool" where
+definition congruent2 :: "('a \<times> 'a) set \<Rightarrow> ('b \<times> 'b) set \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> bool" where
   "congruent2 r1 r2 f \<longleftrightarrow> (\<forall>(y1, z1) \<in> r1. \<forall>(y2, z2) \<in> r2. f y1 y2 = f z1 z2)"
 
 lemma congruent2I':
--- a/src/HOL/Fun.thy	Thu Aug 18 13:36:58 2011 -0700
+++ b/src/HOL/Fun.thy	Thu Aug 18 14:08:39 2011 -0700
@@ -10,15 +10,6 @@
 uses ("Tools/enriched_type.ML")
 begin
 
-text{*As a simplification rule, it replaces all function equalities by
-  first-order equalities.*}
-lemma fun_eq_iff: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)"
-apply (rule iffI)
-apply (simp (no_asm_simp))
-apply (rule ext)
-apply (simp (no_asm_simp))
-done
-
 lemma apply_inverse:
   "f x = u \<Longrightarrow> (\<And>x. P x \<Longrightarrow> g (f x) = x) \<Longrightarrow> P x \<Longrightarrow> x = g u"
   by auto
@@ -26,26 +17,22 @@
 
 subsection {* The Identity Function @{text id} *}
 
-definition
-  id :: "'a \<Rightarrow> 'a"
-where
+definition id :: "'a \<Rightarrow> 'a" where
   "id = (\<lambda>x. x)"
 
 lemma id_apply [simp]: "id x = x"
   by (simp add: id_def)
 
 lemma image_id [simp]: "id ` Y = Y"
-by (simp add: id_def)
+  by (simp add: id_def)
 
 lemma vimage_id [simp]: "id -` A = A"
-by (simp add: id_def)
+  by (simp add: id_def)
 
 
 subsection {* The Composition Operator @{text "f \<circ> g"} *}
 
-definition
-  comp :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "o" 55)
-where
+definition comp :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "o" 55) where
   "f o g = (\<lambda>x. f (g x))"
 
 notation (xsymbols)
@@ -54,9 +41,6 @@
 notation (HTML output)
   comp  (infixl "\<circ>" 55)
 
-text{*compatibility*}
-lemmas o_def = comp_def
-
 lemma o_apply [simp]: "(f o g) x = f (g x)"
 by (simp add: comp_def)
 
@@ -71,7 +55,7 @@
 
 lemma o_eq_dest:
   "a o b = c o d \<Longrightarrow> a (b v) = c (d v)"
-  by (simp only: o_def) (fact fun_cong)
+  by (simp only: comp_def) (fact fun_cong)
 
 lemma o_eq_elim:
   "a o b = c o d \<Longrightarrow> ((\<And>v. a (b v) = c (d v)) \<Longrightarrow> R) \<Longrightarrow> R"
@@ -89,9 +73,7 @@
 
 subsection {* The Forward Composition Operator @{text fcomp} *}
 
-definition
-  fcomp :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "\<circ>>" 60)
-where
+definition fcomp :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "\<circ>>" 60) where
   "f \<circ>> g = (\<lambda>x. g (f x))"
 
 lemma fcomp_apply [simp]:  "(f \<circ>> g) x = g (f x)"
@@ -569,8 +551,7 @@
 
 subsection{*Function Updating*}
 
-definition
-  fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" where
+definition fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" where
   "fun_upd f a b == % x. if x=a then b else f x"
 
 nonterminal updbinds and updbind
@@ -634,9 +615,7 @@
 
 subsection {* @{text override_on} *}
 
-definition
-  override_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b"
-where
+definition override_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b" where
   "override_on f g A = (\<lambda>a. if a \<in> A then g a else f a)"
 
 lemma override_on_emptyset[simp]: "override_on f g {} = f"
@@ -651,9 +630,7 @@
 
 subsection {* @{text swap} *}
 
-definition
-  swap :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)"
-where
+definition swap :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)" where
   "swap a b f = f (a := f b, b:= f a)"
 
 lemma swap_self [simp]: "swap a a f = f"
@@ -716,7 +693,7 @@
 subsection {* Inversion of injective functions *}
 
 definition the_inv_into :: "'a set => ('a => 'b) => ('b => 'a)" where
-"the_inv_into A f == %x. THE y. y : A & f y = x"
+  "the_inv_into A f == %x. THE y. y : A & f y = x"
 
 lemma the_inv_into_f_f:
   "[| inj_on f A;  x : A |] ==> the_inv_into A f (f x) = x"
@@ -775,6 +752,11 @@
   shows "the_inv f (f x) = x" using assms UNIV_I
   by (rule the_inv_into_f_f)
 
+
+text{*compatibility*}
+lemmas o_def = comp_def
+
+
 subsection {* Cantor's Paradox *}
 
 lemma Cantors_paradox [no_atp]:
--- a/src/HOL/GCD.thy	Thu Aug 18 13:36:58 2011 -0700
+++ b/src/HOL/GCD.thy	Thu Aug 18 14:08:39 2011 -0700
@@ -531,11 +531,8 @@
 
 lemma Max_divisors_self_int[simp]: "n\<noteq>0 \<Longrightarrow> Max{d::int. d dvd n} = abs n"
 apply(rule antisym)
- apply(rule Max_le_iff[THEN iffD2])
-   apply simp
-  apply fastsimp
- apply (metis Collect_def abs_ge_self dvd_imp_le_int mem_def zle_trans)
-apply simp
+ apply(rule Max_le_iff [THEN iffD2])
+  apply (auto intro: abs_le_D1 dvd_imp_le_int)
 done
 
 lemma gcd_is_Max_divisors_nat:
--- a/src/HOL/HOL.thy	Thu Aug 18 13:36:58 2011 -0700
+++ b/src/HOL/HOL.thy	Thu Aug 18 14:08:39 2011 -0700
@@ -1394,6 +1394,11 @@
   "f (if c then x else y) = (if c then f x else f y)"
   by simp
 
+text{*As a simplification rule, it replaces all function equalities by
+  first-order equalities.*}
+lemma fun_eq_iff: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)"
+  by auto
+
 
 subsubsection {* Generic cases and induction *}
 
--- a/src/HOL/IsaMakefile	Thu Aug 18 13:36:58 2011 -0700
+++ b/src/HOL/IsaMakefile	Thu Aug 18 14:08:39 2011 -0700
@@ -1039,32 +1039,30 @@
 
 $(LOG)/HOL-ex.gz: $(OUT)/HOL Decision_Procs/Commutative_Ring.thy	\
   Number_Theory/Primes.thy ex/Abstract_NAT.thy ex/Antiquote.thy		\
-  ex/Arith_Examples.thy ex/Arithmetic_Series_Complex.thy \
-  ex/BT.thy	ex/BinEx.thy ex/Binary.thy ex/Birthday_Paradox.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		\
-  ex/Efficient_Nat_examples.thy ex/Eval_Examples.thy ex/Fundefs.thy	\
-  ex/Gauge_Integration.thy ex/Groebner_Examples.thy ex/Guess.thy	\
-  ex/HarmonicSeries.thy ex/Hebrew.thy ex/Hex_Bin_Examples.thy		\
-  ex/Higher_Order_Logic.thy ex/Iff_Oracle.thy ex/Induction_Schema.thy	\
+  ex/Arith_Examples.thy ex/Arithmetic_Series_Complex.thy ex/BT.thy	\
+  ex/BinEx.thy ex/Binary.thy ex/Birthday_Paradox.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 ex/Efficient_Nat_examples.thy			\
+  ex/Eval_Examples.thy ex/Fundefs.thy ex/Gauge_Integration.thy		\
+  ex/Groebner_Examples.thy ex/Guess.thy ex/HarmonicSeries.thy		\
+  ex/Hebrew.thy ex/Hex_Bin_Examples.thy ex/Higher_Order_Logic.thy	\
+  ex/Iff_Oracle.thy ex/Induction_Schema.thy				\
   ex/Interpretation_with_Defs.thy ex/Intuitionistic.thy ex/Lagrange.thy	\
   ex/List_to_Set_Comprehension_Examples.thy ex/LocaleTest2.thy		\
   ex/MT.thy ex/MergeSort.thy ex/Meson_Test.thy ex/MonoidGroup.thy	\
   ex/Multiquote.thy ex/NatSum.thy ex/Normalization_by_Evaluation.thy	\
   ex/Numeral.thy ex/PER.thy ex/PresburgerEx.thy ex/Primrec.thy		\
   ex/Quickcheck_Examples.thy ex/Quickcheck_Lattice_Examples.thy		\
-  ex/Quickcheck_Narrowing_Examples.thy  				\
-  ex/Quicksort.thy ex/ROOT.ML ex/Records.thy		\
-  ex/ReflectionEx.thy ex/Refute_Examples.thy ex/SAT_Examples.thy	\
-  ex/SVC_Oracle.thy ex/Serbian.thy ex/Set_Algebras.thy			\
-  ex/sledgehammer_tactics.ML ex/Sqrt.thy				\
-  ex/Sqrt_Script.thy ex/Sudoku.thy ex/Tarski.thy ex/Termination.thy	\
-  ex/Transfer_Ex.thy ex/Tree23.thy			\
+  ex/Quickcheck_Narrowing_Examples.thy ex/Quicksort.thy ex/ROOT.ML	\
+  ex/Records.thy ex/ReflectionEx.thy ex/Refute_Examples.thy		\
+  ex/SAT_Examples.thy ex/Serbian.thy ex/Set_Theory.thy			\
+  ex/Set_Algebras.thy ex/SVC_Oracle.thy ex/sledgehammer_tactics.ML	\
+  ex/Sqrt.thy ex/Sqrt_Script.thy ex/Sudoku.thy ex/Tarski.thy		\
+  ex/Termination.thy ex/Transfer_Ex.thy ex/Tree23.thy			\
   ex/Unification.thy ex/While_Combinator_Example.thy			\
-  ex/document/root.bib ex/document/root.tex ex/set.thy ex/svc_funcs.ML	\
-  ex/svc_test.thy							\
-  ../Tools/interpretation_with_defs.ML
+  ex/document/root.bib ex/document/root.tex ex/svc_funcs.ML		\
+  ex/svc_test.thy ../Tools/interpretation_with_defs.ML
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL ex
 
 
--- a/src/HOL/Nat.thy	Thu Aug 18 13:36:58 2011 -0700
+++ b/src/HOL/Nat.thy	Thu Aug 18 14:08:39 2011 -0700
@@ -39,11 +39,20 @@
     Zero_RepI: "Nat Zero_Rep"
   | Suc_RepI: "Nat i \<Longrightarrow> Nat (Suc_Rep i)"
 
-typedef (open Nat) nat = Nat
-  by (rule exI, unfold mem_def, rule Nat.Zero_RepI)
+typedef (open Nat) nat = "{n. Nat n}"
+  using Nat.Zero_RepI by auto
+
+lemma Nat_Rep_Nat:
+  "Nat (Rep_Nat n)"
+  using Rep_Nat by simp
 
-definition Suc :: "nat => nat" where
-  "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))"
+lemma Nat_Abs_Nat_inverse:
+  "Nat n \<Longrightarrow> Rep_Nat (Abs_Nat n) = n"
+  using Abs_Nat_inverse by simp
+
+lemma Nat_Abs_Nat_inject:
+  "Nat n \<Longrightarrow> Nat m \<Longrightarrow> Abs_Nat n = Abs_Nat m \<longleftrightarrow> n = m"
+  using Abs_Nat_inject by simp
 
 instantiation nat :: zero
 begin
@@ -55,9 +64,11 @@
 
 end
 
+definition Suc :: "nat \<Rightarrow> nat" where
+  "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))"
+
 lemma Suc_not_Zero: "Suc m \<noteq> 0"
-  by (simp add: Zero_nat_def Suc_def Abs_Nat_inject [unfolded mem_def]
-    Rep_Nat [unfolded mem_def] Suc_RepI Zero_RepI Suc_Rep_not_Zero_Rep [unfolded mem_def])
+  by (simp add: Zero_nat_def Suc_def Suc_RepI Zero_RepI Nat_Abs_Nat_inject Suc_Rep_not_Zero_Rep Nat_Rep_Nat)
 
 lemma Zero_not_Suc: "0 \<noteq> Suc m"
   by (rule not_sym, rule Suc_not_Zero not_sym)
@@ -67,12 +78,12 @@
 
 rep_datatype "0 \<Colon> nat" Suc
   apply (unfold Zero_nat_def Suc_def)
-     apply (rule Rep_Nat_inverse [THEN subst]) -- {* types force good instantiation *}
-     apply (erule Rep_Nat [unfolded mem_def, THEN Nat.induct])
-     apply (iprover elim: Abs_Nat_inverse [unfolded mem_def, THEN subst])
-    apply (simp_all add: Abs_Nat_inject [unfolded mem_def] Rep_Nat [unfolded mem_def]
-      Suc_RepI Zero_RepI Suc_Rep_not_Zero_Rep [unfolded mem_def]
-      Suc_Rep_not_Zero_Rep [unfolded mem_def, symmetric]
+  apply (rule Rep_Nat_inverse [THEN subst]) -- {* types force good instantiation *}
+   apply (erule Nat_Rep_Nat [THEN Nat.induct])
+   apply (iprover elim: Nat_Abs_Nat_inverse [THEN subst])
+    apply (simp_all add: Nat_Abs_Nat_inject Nat_Rep_Nat
+      Suc_RepI Zero_RepI Suc_Rep_not_Zero_Rep
+      Suc_Rep_not_Zero_Rep [symmetric]
       Suc_Rep_inject' Rep_Nat_inject)
   done
 
--- a/src/HOL/Nitpick.thy	Thu Aug 18 13:36:58 2011 -0700
+++ b/src/HOL/Nitpick.thy	Thu Aug 18 14:08:39 2011 -0700
@@ -76,19 +76,19 @@
 definition prod :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set" where
 "prod A B = {(a, b). a \<in> A \<and> b \<in> B}"
 
-definition refl' :: "('a \<times> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
+definition refl' :: "('a \<times> 'a) set \<Rightarrow> bool" where
 "refl' r \<equiv> \<forall>x. (x, x) \<in> r"
 
-definition wf' :: "('a \<times> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
+definition wf' :: "('a \<times> 'a) set \<Rightarrow> bool" where
 "wf' r \<equiv> acyclic r \<and> (finite r \<or> unknown)"
 
-definition card' :: "('a \<Rightarrow> bool) \<Rightarrow> nat" where
+definition card' :: "'a set \<Rightarrow> nat" where
 "card' A \<equiv> if finite A then length (SOME xs. set xs = A \<and> distinct xs) else 0"
 
-definition setsum' :: "('a \<Rightarrow> 'b\<Colon>comm_monoid_add) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'b" where
+definition setsum' :: "('a \<Rightarrow> 'b\<Colon>comm_monoid_add) \<Rightarrow> 'a set \<Rightarrow> 'b" where
 "setsum' f A \<equiv> if finite A then listsum (map f (SOME xs. set xs = A \<and> distinct xs)) else 0"
 
-inductive fold_graph' :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> bool" where
+inductive fold_graph' :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool" where
 "fold_graph' f z {} z" |
 "\<lbrakk>x \<in> A; fold_graph' f z (A - {x}) y\<rbrakk> \<Longrightarrow> fold_graph' f z A (f x y)"
 
--- a/src/HOL/RealDef.thy	Thu Aug 18 13:36:58 2011 -0700
+++ b/src/HOL/RealDef.thy	Thu Aug 18 14:08:39 2011 -0700
@@ -121,7 +121,7 @@
 subsection {* Cauchy sequences *}
 
 definition
-  cauchy :: "(nat \<Rightarrow> rat) set"
+  cauchy :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
 where
   "cauchy X \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r)"
 
--- a/src/HOL/Relation.thy	Thu Aug 18 13:36:58 2011 -0700
+++ b/src/HOL/Relation.thy	Thu Aug 18 14:08:39 2011 -0700
@@ -133,9 +133,8 @@
 by blast
 
 lemma Id_on_def' [nitpick_unfold, code]:
-  "(Id_on (A :: 'a => bool)) = (%(x, y). x = y \<and> A x)"
-by (auto simp add: fun_eq_iff
-  elim: Id_onE[unfolded mem_def] intro: Id_onI[unfolded mem_def])
+  "Id_on {x. A x} = Collect (\<lambda>(x, y). x = y \<and> A x)"
+by auto
 
 lemma Id_on_subset_Times: "Id_on A \<subseteq> A \<times> A"
 by blast
--- a/src/HOL/String.thy	Thu Aug 18 13:36:58 2011 -0700
+++ b/src/HOL/String.thy	Thu Aug 18 14:08:39 2011 -0700
@@ -155,7 +155,7 @@
 
 subsection {* Strings as dedicated type *}
 
-typedef (open) literal = "UNIV :: string \<Rightarrow> bool"
+typedef (open) literal = "UNIV :: string set"
   morphisms explode STR ..
 
 instantiation literal :: size
--- a/src/HOL/ex/ROOT.ML	Thu Aug 18 13:36:58 2011 -0700
+++ b/src/HOL/ex/ROOT.ML	Thu Aug 18 14:08:39 2011 -0700
@@ -48,7 +48,7 @@
   "Primrec",
   "Tarski",
   "Classical",
-  "set",
+  "Set_Theory",
   "Meson_Test",
   "Termination",
   "Coherent",
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Set_Theory.thy	Thu Aug 18 14:08:39 2011 -0700
@@ -0,0 +1,227 @@
+(*  Title:      HOL/ex/Set_Theory.thy
+    Author:     Tobias Nipkow and Lawrence C Paulson
+    Copyright   1991  University of Cambridge
+*)
+
+header {* Set Theory examples: Cantor's Theorem, Schröder-Bernstein Theorem, etc. *}
+
+theory Set_Theory
+imports Main
+begin
+
+text{*
+  These two are cited in Benzmueller and Kohlhase's system description
+  of LEO, CADE-15, 1998 (pages 139-143) as theorems LEO could not
+  prove.
+*}
+
+lemma "(X = Y \<union> Z) =
+    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
+  by blast
+
+lemma "(X = Y \<inter> Z) =
+    (X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))"
+  by blast
+
+text {*
+  Trivial example of term synthesis: apparently hard for some provers!
+*}
+
+schematic_lemma "a \<noteq> b \<Longrightarrow> a \<in> ?X \<and> b \<notin> ?X"
+  by blast
+
+
+subsection {* Examples for the @{text blast} paper *}
+
+lemma "(\<Union>x \<in> C. f x \<union> g x) = \<Union>(f ` C)  \<union>  \<Union>(g ` C)"
+  -- {* Union-image, called @{text Un_Union_image} in Main HOL *}
+  by blast
+
+lemma "(\<Inter>x \<in> C. f x \<inter> g x) = \<Inter>(f ` C) \<inter> \<Inter>(g ` C)"
+  -- {* Inter-image, called @{text Int_Inter_image} in Main HOL *}
+  by blast
+
+lemma singleton_example_1:
+     "\<And>S::'a set set. \<forall>x \<in> S. \<forall>y \<in> S. x \<subseteq> y \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
+  by blast
+
+lemma singleton_example_2:
+     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
+  -- {*Variant of the problem above. *}
+  by blast
+
+lemma "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"
+  -- {* A unique fixpoint theorem --- @{text fast}/@{text best}/@{text meson} all fail. *}
+  by metis
+
+
+subsection {* Cantor's Theorem: There is no surjection from a set to its powerset *}
+
+lemma cantor1: "\<not> (\<exists>f:: 'a \<Rightarrow> 'a set. \<forall>S. \<exists>x. f x = S)"
+  -- {* Requires best-first search because it is undirectional. *}
+  by best
+
+schematic_lemma "\<forall>f:: 'a \<Rightarrow> 'a set. \<forall>x. f x \<noteq> ?S f"
+  -- {*This form displays the diagonal term. *}
+  by best
+
+schematic_lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
+  -- {* This form exploits the set constructs. *}
+  by (rule notI, erule rangeE, best)
+
+schematic_lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
+  -- {* Or just this! *}
+  by best
+
+
+subsection {* The Schröder-Berstein Theorem *}
+
+lemma disj_lemma: "- (f ` X) = g ` (-X) \<Longrightarrow> f a = g b \<Longrightarrow> a \<in> X \<Longrightarrow> b \<in> X"
+  by blast
+
+lemma surj_if_then_else:
+  "-(f ` X) = g ` (-X) \<Longrightarrow> surj (\<lambda>z. if z \<in> X then f z else g z)"
+  by (simp add: surj_def) blast
+
+lemma bij_if_then_else:
+  "inj_on f X \<Longrightarrow> inj_on g (-X) \<Longrightarrow> -(f ` X) = g ` (-X) \<Longrightarrow>
+    h = (\<lambda>z. if z \<in> X then f z else g z) \<Longrightarrow> inj h \<and> surj h"
+  apply (unfold inj_on_def)
+  apply (simp add: surj_if_then_else)
+  apply (blast dest: disj_lemma sym)
+  done
+
+lemma decomposition: "\<exists>X. X = - (g ` (- (f ` X)))"
+  apply (rule exI)
+  apply (rule lfp_unfold)
+  apply (rule monoI, blast)
+  done
+
+theorem Schroeder_Bernstein:
+  "inj (f :: 'a \<Rightarrow> 'b) \<Longrightarrow> inj (g :: 'b \<Rightarrow> 'a)
+    \<Longrightarrow> \<exists>h:: 'a \<Rightarrow> 'b. inj h \<and> surj h"
+  apply (rule decomposition [where f=f and g=g, THEN exE])
+  apply (rule_tac x = "(\<lambda>z. if z \<in> x then f z else inv g z)" in exI) 
+    --{*The term above can be synthesized by a sufficiently detailed proof.*}
+  apply (rule bij_if_then_else)
+     apply (rule_tac [4] refl)
+    apply (rule_tac [2] inj_on_inv_into)
+    apply (erule subset_inj_on [OF _ subset_UNIV])
+   apply blast
+  apply (erule ssubst, subst double_complement, erule inv_image_comp [symmetric])
+  done
+
+
+subsection {* A simple party theorem *}
+
+text{* \emph{At any party there are two people who know the same
+number of people}. Provided the party consists of at least two people
+and the knows relation is symmetric. Knowing yourself does not count
+--- otherwise knows needs to be reflexive. (From Freek Wiedijk's talk
+at TPHOLs 2007.) *}
+
+lemma equal_number_of_acquaintances:
+assumes "Domain R <= A" and "sym R" and "card A \<ge> 2"
+shows "\<not> inj_on (%a. card(R `` {a} - {a})) A"
+proof -
+  let ?N = "%a. card(R `` {a} - {a})"
+  let ?n = "card A"
+  have "finite A" using `card A \<ge> 2` by(auto intro:ccontr)
+  have 0: "R `` A <= A" using `sym R` `Domain R <= A`
+    unfolding Domain_def sym_def by blast
+  have h: "ALL a:A. R `` {a} <= A" using 0 by blast
+  hence 1: "ALL a:A. finite(R `` {a})" using `finite A`
+    by(blast intro: finite_subset)
+  have sub: "?N ` A <= {0..<?n}"
+  proof -
+    have "ALL a:A. R `` {a} - {a} < A" using h by blast
+    thus ?thesis using psubset_card_mono[OF `finite A`] by auto
+  qed
+  show "~ inj_on ?N A" (is "~ ?I")
+  proof
+    assume ?I
+    hence "?n = card(?N ` A)" by(rule card_image[symmetric])
+    with sub `finite A` have 2[simp]: "?N ` A = {0..<?n}"
+      using subset_card_intvl_is_intvl[of _ 0] by(auto)
+    have "0 : ?N ` A" and "?n - 1 : ?N ` A"  using `card A \<ge> 2` by simp+
+    then obtain a b where ab: "a:A" "b:A" and Na: "?N a = 0" and Nb: "?N b = ?n - 1"
+      by (auto simp del: 2)
+    have "a \<noteq> b" using Na Nb `card A \<ge> 2` by auto
+    have "R `` {a} - {a} = {}" by (metis 1 Na ab card_eq_0_iff finite_Diff)
+    hence "b \<notin> R `` {a}" using `a\<noteq>b` by blast
+    hence "a \<notin> R `` {b}" by (metis Image_singleton_iff assms(2) sym_def)
+    hence 3: "R `` {b} - {b} <= A - {a,b}" using 0 ab by blast
+    have 4: "finite (A - {a,b})" using `finite A` by simp
+    have "?N b <= ?n - 2" using ab `a\<noteq>b` `finite A` card_mono[OF 4 3] by simp
+    then show False using Nb `card A \<ge>  2` by arith
+  qed
+qed
+
+text {*
+  From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages
+  293-314.
+
+  Isabelle can prove the easy examples without any special mechanisms,
+  but it can't prove the hard ones.
+*}
+
+lemma "\<exists>A. (\<forall>x \<in> A. x \<le> (0::int))"
+  -- {* Example 1, page 295. *}
+  by force
+
+lemma "D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B"
+  -- {* Example 2. *}
+  by force
+
+lemma "P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)"
+  -- {* Example 3. *}
+  by force
+
+lemma "a < b \<and> b < (c::int) \<Longrightarrow> \<exists>A. a \<notin> A \<and> b \<in> A \<and> c \<notin> A"
+  -- {* Example 4. *}
+  by force
+
+lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
+  -- {*Example 5, page 298. *}
+  by force
+
+lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
+  -- {* Example 6. *}
+  by force
+
+lemma "\<exists>A. a \<notin> A"
+  -- {* Example 7. *}
+  by force
+
+lemma "(\<forall>u v. u < (0::int) \<longrightarrow> u \<noteq> abs v)
+    \<longrightarrow> (\<exists>A::int set. (\<forall>y. abs y \<notin> A) \<and> -2 \<in> A)"
+  -- {* Example 8 now needs a small hint. *}
+  by (simp add: abs_if, force)
+    -- {* not @{text blast}, which can't simplify @{text "-2 < 0"} *}
+
+text {* Example 9 omitted (requires the reals). *}
+
+text {* The paper has no Example 10! *}
+
+lemma "(\<forall>A. 0 \<in> A \<and> (\<forall>x \<in> A. Suc x \<in> A) \<longrightarrow> n \<in> A) \<and>
+  P 0 \<and> (\<forall>x. P x \<longrightarrow> P (Suc x)) \<longrightarrow> P n"
+  -- {* Example 11: needs a hint. *}
+by(metis nat.induct)
+
+lemma
+  "(\<forall>A. (0, 0) \<in> A \<and> (\<forall>x y. (x, y) \<in> A \<longrightarrow> (Suc x, Suc y) \<in> A) \<longrightarrow> (n, m) \<in> A)
+    \<and> P n \<longrightarrow> P m"
+  -- {* Example 12. *}
+  by auto
+
+lemma
+  "(\<forall>x. (\<exists>u. x = 2 * u) = (\<not> (\<exists>v. Suc x = 2 * v))) \<longrightarrow>
+    (\<exists>A. \<forall>x. (x \<in> A) = (Suc x \<notin> A))"
+  -- {* Example EO1: typo in article, and with the obvious fix it seems
+      to require arithmetic reasoning. *}
+  apply clarify
+  apply (rule_tac x = "{x. \<exists>u. x = 2 * u}" in exI, auto)
+   apply metis+
+  done
+
+end
--- a/src/HOL/ex/set.thy	Thu Aug 18 13:36:58 2011 -0700
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,225 +0,0 @@
-(*  Title:      HOL/ex/set.thy
-    Author:     Tobias Nipkow and Lawrence C Paulson
-    Copyright   1991  University of Cambridge
-*)
-
-header {* Set Theory examples: Cantor's Theorem, Schröder-Bernstein Theorem, etc. *}
-
-theory set imports Main begin
-
-text{*
-  These two are cited in Benzmueller and Kohlhase's system description
-  of LEO, CADE-15, 1998 (pages 139-143) as theorems LEO could not
-  prove.
-*}
-
-lemma "(X = Y \<union> Z) =
-    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
-  by blast
-
-lemma "(X = Y \<inter> Z) =
-    (X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))"
-  by blast
-
-text {*
-  Trivial example of term synthesis: apparently hard for some provers!
-*}
-
-schematic_lemma "a \<noteq> b \<Longrightarrow> a \<in> ?X \<and> b \<notin> ?X"
-  by blast
-
-
-subsection {* Examples for the @{text blast} paper *}
-
-lemma "(\<Union>x \<in> C. f x \<union> g x) = \<Union>(f ` C)  \<union>  \<Union>(g ` C)"
-  -- {* Union-image, called @{text Un_Union_image} in Main HOL *}
-  by blast
-
-lemma "(\<Inter>x \<in> C. f x \<inter> g x) = \<Inter>(f ` C) \<inter> \<Inter>(g ` C)"
-  -- {* Inter-image, called @{text Int_Inter_image} in Main HOL *}
-  by blast
-
-lemma singleton_example_1:
-     "\<And>S::'a set set. \<forall>x \<in> S. \<forall>y \<in> S. x \<subseteq> y \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
-  by blast
-
-lemma singleton_example_2:
-     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
-  -- {*Variant of the problem above. *}
-  by blast
-
-lemma "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"
-  -- {* A unique fixpoint theorem --- @{text fast}/@{text best}/@{text meson} all fail. *}
-  by metis
-
-
-subsection {* Cantor's Theorem: There is no surjection from a set to its powerset *}
-
-lemma cantor1: "\<not> (\<exists>f:: 'a \<Rightarrow> 'a set. \<forall>S. \<exists>x. f x = S)"
-  -- {* Requires best-first search because it is undirectional. *}
-  by best
-
-schematic_lemma "\<forall>f:: 'a \<Rightarrow> 'a set. \<forall>x. f x \<noteq> ?S f"
-  -- {*This form displays the diagonal term. *}
-  by best
-
-schematic_lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
-  -- {* This form exploits the set constructs. *}
-  by (rule notI, erule rangeE, best)
-
-schematic_lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
-  -- {* Or just this! *}
-  by best
-
-
-subsection {* The Schröder-Berstein Theorem *}
-
-lemma disj_lemma: "- (f ` X) = g ` (-X) \<Longrightarrow> f a = g b \<Longrightarrow> a \<in> X \<Longrightarrow> b \<in> X"
-  by blast
-
-lemma surj_if_then_else:
-  "-(f ` X) = g ` (-X) \<Longrightarrow> surj (\<lambda>z. if z \<in> X then f z else g z)"
-  by (simp add: surj_def) blast
-
-lemma bij_if_then_else:
-  "inj_on f X \<Longrightarrow> inj_on g (-X) \<Longrightarrow> -(f ` X) = g ` (-X) \<Longrightarrow>
-    h = (\<lambda>z. if z \<in> X then f z else g z) \<Longrightarrow> inj h \<and> surj h"
-  apply (unfold inj_on_def)
-  apply (simp add: surj_if_then_else)
-  apply (blast dest: disj_lemma sym)
-  done
-
-lemma decomposition: "\<exists>X. X = - (g ` (- (f ` X)))"
-  apply (rule exI)
-  apply (rule lfp_unfold)
-  apply (rule monoI, blast)
-  done
-
-theorem Schroeder_Bernstein:
-  "inj (f :: 'a \<Rightarrow> 'b) \<Longrightarrow> inj (g :: 'b \<Rightarrow> 'a)
-    \<Longrightarrow> \<exists>h:: 'a \<Rightarrow> 'b. inj h \<and> surj h"
-  apply (rule decomposition [where f=f and g=g, THEN exE])
-  apply (rule_tac x = "(\<lambda>z. if z \<in> x then f z else inv g z)" in exI) 
-    --{*The term above can be synthesized by a sufficiently detailed proof.*}
-  apply (rule bij_if_then_else)
-     apply (rule_tac [4] refl)
-    apply (rule_tac [2] inj_on_inv_into)
-    apply (erule subset_inj_on [OF _ subset_UNIV])
-   apply blast
-  apply (erule ssubst, subst double_complement, erule inv_image_comp [symmetric])
-  done
-
-
-subsection {* A simple party theorem *}
-
-text{* \emph{At any party there are two people who know the same
-number of people}. Provided the party consists of at least two people
-and the knows relation is symmetric. Knowing yourself does not count
---- otherwise knows needs to be reflexive. (From Freek Wiedijk's talk
-at TPHOLs 2007.) *}
-
-lemma equal_number_of_acquaintances:
-assumes "Domain R <= A" and "sym R" and "card A \<ge> 2"
-shows "\<not> inj_on (%a. card(R `` {a} - {a})) A"
-proof -
-  let ?N = "%a. card(R `` {a} - {a})"
-  let ?n = "card A"
-  have "finite A" using `card A \<ge> 2` by(auto intro:ccontr)
-  have 0: "R `` A <= A" using `sym R` `Domain R <= A`
-    unfolding Domain_def sym_def by blast
-  have h: "ALL a:A. R `` {a} <= A" using 0 by blast
-  hence 1: "ALL a:A. finite(R `` {a})" using `finite A`
-    by(blast intro: finite_subset)
-  have sub: "?N ` A <= {0..<?n}"
-  proof -
-    have "ALL a:A. R `` {a} - {a} < A" using h by blast
-    thus ?thesis using psubset_card_mono[OF `finite A`] by auto
-  qed
-  show "~ inj_on ?N A" (is "~ ?I")
-  proof
-    assume ?I
-    hence "?n = card(?N ` A)" by(rule card_image[symmetric])
-    with sub `finite A` have 2[simp]: "?N ` A = {0..<?n}"
-      using subset_card_intvl_is_intvl[of _ 0] by(auto)
-    have "0 : ?N ` A" and "?n - 1 : ?N ` A"  using `card A \<ge> 2` by simp+
-    then obtain a b where ab: "a:A" "b:A" and Na: "?N a = 0" and Nb: "?N b = ?n - 1"
-      by (auto simp del: 2)
-    have "a \<noteq> b" using Na Nb `card A \<ge> 2` by auto
-    have "R `` {a} - {a} = {}" by (metis 1 Na ab card_eq_0_iff finite_Diff)
-    hence "b \<notin> R `` {a}" using `a\<noteq>b` by blast
-    hence "a \<notin> R `` {b}" by (metis Image_singleton_iff assms(2) sym_def)
-    hence 3: "R `` {b} - {b} <= A - {a,b}" using 0 ab by blast
-    have 4: "finite (A - {a,b})" using `finite A` by simp
-    have "?N b <= ?n - 2" using ab `a\<noteq>b` `finite A` card_mono[OF 4 3] by simp
-    then show False using Nb `card A \<ge>  2` by arith
-  qed
-qed
-
-text {*
-  From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages
-  293-314.
-
-  Isabelle can prove the easy examples without any special mechanisms,
-  but it can't prove the hard ones.
-*}
-
-lemma "\<exists>A. (\<forall>x \<in> A. x \<le> (0::int))"
-  -- {* Example 1, page 295. *}
-  by force
-
-lemma "D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B"
-  -- {* Example 2. *}
-  by force
-
-lemma "P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)"
-  -- {* Example 3. *}
-  by force
-
-lemma "a < b \<and> b < (c::int) \<Longrightarrow> \<exists>A. a \<notin> A \<and> b \<in> A \<and> c \<notin> A"
-  -- {* Example 4. *}
-  by force
-
-lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
-  -- {*Example 5, page 298. *}
-  by force
-
-lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
-  -- {* Example 6. *}
-  by force
-
-lemma "\<exists>A. a \<notin> A"
-  -- {* Example 7. *}
-  by force
-
-lemma "(\<forall>u v. u < (0::int) \<longrightarrow> u \<noteq> abs v)
-    \<longrightarrow> (\<exists>A::int set. (\<forall>y. abs y \<notin> A) \<and> -2 \<in> A)"
-  -- {* Example 8 now needs a small hint. *}
-  by (simp add: abs_if, force)
-    -- {* not @{text blast}, which can't simplify @{text "-2 < 0"} *}
-
-text {* Example 9 omitted (requires the reals). *}
-
-text {* The paper has no Example 10! *}
-
-lemma "(\<forall>A. 0 \<in> A \<and> (\<forall>x \<in> A. Suc x \<in> A) \<longrightarrow> n \<in> A) \<and>
-  P 0 \<and> (\<forall>x. P x \<longrightarrow> P (Suc x)) \<longrightarrow> P n"
-  -- {* Example 11: needs a hint. *}
-by(metis nat.induct)
-
-lemma
-  "(\<forall>A. (0, 0) \<in> A \<and> (\<forall>x y. (x, y) \<in> A \<longrightarrow> (Suc x, Suc y) \<in> A) \<longrightarrow> (n, m) \<in> A)
-    \<and> P n \<longrightarrow> P m"
-  -- {* Example 12. *}
-  by auto
-
-lemma
-  "(\<forall>x. (\<exists>u. x = 2 * u) = (\<not> (\<exists>v. Suc x = 2 * v))) \<longrightarrow>
-    (\<exists>A. \<forall>x. (x \<in> A) = (Suc x \<notin> A))"
-  -- {* Example EO1: typo in article, and with the obvious fix it seems
-      to require arithmetic reasoning. *}
-  apply clarify
-  apply (rule_tac x = "{x. \<exists>u. x = 2 * u}" in exI, auto)
-   apply metis+
-  done
-
-end