merged
authorhaftmann
Thu, 18 Aug 2011 22:50:28 +0200
changeset 45144d27b9fe4759e
parent 45138 d995733b635d
parent 45143 4846f3f320d9
child 45146 637297cf6142
merged
src/HOL/ex/set.thy
     1.1 --- a/src/HOL/Equiv_Relations.thy	Thu Aug 18 17:42:35 2011 +0200
     1.2 +++ b/src/HOL/Equiv_Relations.thy	Thu Aug 18 22:50:28 2011 +0200
     1.3 @@ -164,7 +164,7 @@
     1.4  
     1.5  text{*A congruence-preserving function*}
     1.6  
     1.7 -definition congruent :: "('a \<times> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"  where
     1.8 +definition congruent :: "('a \<times> 'a) set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"  where
     1.9    "congruent r f \<longleftrightarrow> (\<forall>(y, z) \<in> r. f y = f z)"
    1.10  
    1.11  lemma congruentI:
    1.12 @@ -229,7 +229,7 @@
    1.13  
    1.14  text{*A congruence-preserving function of two arguments*}
    1.15  
    1.16 -definition congruent2 :: "('a \<times> 'a \<Rightarrow> bool) \<Rightarrow> ('b \<times> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> bool" where
    1.17 +definition congruent2 :: "('a \<times> 'a) set \<Rightarrow> ('b \<times> 'b) set \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> bool" where
    1.18    "congruent2 r1 r2 f \<longleftrightarrow> (\<forall>(y1, z1) \<in> r1. \<forall>(y2, z2) \<in> r2. f y1 y2 = f z1 z2)"
    1.19  
    1.20  lemma congruent2I':
     2.1 --- a/src/HOL/Fun.thy	Thu Aug 18 17:42:35 2011 +0200
     2.2 +++ b/src/HOL/Fun.thy	Thu Aug 18 22:50:28 2011 +0200
     2.3 @@ -10,15 +10,6 @@
     2.4  uses ("Tools/enriched_type.ML")
     2.5  begin
     2.6  
     2.7 -text{*As a simplification rule, it replaces all function equalities by
     2.8 -  first-order equalities.*}
     2.9 -lemma fun_eq_iff: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)"
    2.10 -apply (rule iffI)
    2.11 -apply (simp (no_asm_simp))
    2.12 -apply (rule ext)
    2.13 -apply (simp (no_asm_simp))
    2.14 -done
    2.15 -
    2.16  lemma apply_inverse:
    2.17    "f x = u \<Longrightarrow> (\<And>x. P x \<Longrightarrow> g (f x) = x) \<Longrightarrow> P x \<Longrightarrow> x = g u"
    2.18    by auto
    2.19 @@ -26,26 +17,22 @@
    2.20  
    2.21  subsection {* The Identity Function @{text id} *}
    2.22  
    2.23 -definition
    2.24 -  id :: "'a \<Rightarrow> 'a"
    2.25 -where
    2.26 +definition id :: "'a \<Rightarrow> 'a" where
    2.27    "id = (\<lambda>x. x)"
    2.28  
    2.29  lemma id_apply [simp]: "id x = x"
    2.30    by (simp add: id_def)
    2.31  
    2.32  lemma image_id [simp]: "id ` Y = Y"
    2.33 -by (simp add: id_def)
    2.34 +  by (simp add: id_def)
    2.35  
    2.36  lemma vimage_id [simp]: "id -` A = A"
    2.37 -by (simp add: id_def)
    2.38 +  by (simp add: id_def)
    2.39  
    2.40  
    2.41  subsection {* The Composition Operator @{text "f \<circ> g"} *}
    2.42  
    2.43 -definition
    2.44 -  comp :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "o" 55)
    2.45 -where
    2.46 +definition comp :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "o" 55) where
    2.47    "f o g = (\<lambda>x. f (g x))"
    2.48  
    2.49  notation (xsymbols)
    2.50 @@ -54,9 +41,6 @@
    2.51  notation (HTML output)
    2.52    comp  (infixl "\<circ>" 55)
    2.53  
    2.54 -text{*compatibility*}
    2.55 -lemmas o_def = comp_def
    2.56 -
    2.57  lemma o_apply [simp]: "(f o g) x = f (g x)"
    2.58  by (simp add: comp_def)
    2.59  
    2.60 @@ -71,7 +55,7 @@
    2.61  
    2.62  lemma o_eq_dest:
    2.63    "a o b = c o d \<Longrightarrow> a (b v) = c (d v)"
    2.64 -  by (simp only: o_def) (fact fun_cong)
    2.65 +  by (simp only: comp_def) (fact fun_cong)
    2.66  
    2.67  lemma o_eq_elim:
    2.68    "a o b = c o d \<Longrightarrow> ((\<And>v. a (b v) = c (d v)) \<Longrightarrow> R) \<Longrightarrow> R"
    2.69 @@ -89,9 +73,7 @@
    2.70  
    2.71  subsection {* The Forward Composition Operator @{text fcomp} *}
    2.72  
    2.73 -definition
    2.74 -  fcomp :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "\<circ>>" 60)
    2.75 -where
    2.76 +definition fcomp :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "\<circ>>" 60) where
    2.77    "f \<circ>> g = (\<lambda>x. g (f x))"
    2.78  
    2.79  lemma fcomp_apply [simp]:  "(f \<circ>> g) x = g (f x)"
    2.80 @@ -569,8 +551,7 @@
    2.81  
    2.82  subsection{*Function Updating*}
    2.83  
    2.84 -definition
    2.85 -  fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" where
    2.86 +definition fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" where
    2.87    "fun_upd f a b == % x. if x=a then b else f x"
    2.88  
    2.89  nonterminal updbinds and updbind
    2.90 @@ -634,9 +615,7 @@
    2.91  
    2.92  subsection {* @{text override_on} *}
    2.93  
    2.94 -definition
    2.95 -  override_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b"
    2.96 -where
    2.97 +definition override_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b" where
    2.98    "override_on f g A = (\<lambda>a. if a \<in> A then g a else f a)"
    2.99  
   2.100  lemma override_on_emptyset[simp]: "override_on f g {} = f"
   2.101 @@ -651,9 +630,7 @@
   2.102  
   2.103  subsection {* @{text swap} *}
   2.104  
   2.105 -definition
   2.106 -  swap :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)"
   2.107 -where
   2.108 +definition swap :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)" where
   2.109    "swap a b f = f (a := f b, b:= f a)"
   2.110  
   2.111  lemma swap_self [simp]: "swap a a f = f"
   2.112 @@ -716,7 +693,7 @@
   2.113  subsection {* Inversion of injective functions *}
   2.114  
   2.115  definition the_inv_into :: "'a set => ('a => 'b) => ('b => 'a)" where
   2.116 -"the_inv_into A f == %x. THE y. y : A & f y = x"
   2.117 +  "the_inv_into A f == %x. THE y. y : A & f y = x"
   2.118  
   2.119  lemma the_inv_into_f_f:
   2.120    "[| inj_on f A;  x : A |] ==> the_inv_into A f (f x) = x"
   2.121 @@ -775,6 +752,11 @@
   2.122    shows "the_inv f (f x) = x" using assms UNIV_I
   2.123    by (rule the_inv_into_f_f)
   2.124  
   2.125 +
   2.126 +text{*compatibility*}
   2.127 +lemmas o_def = comp_def
   2.128 +
   2.129 +
   2.130  subsection {* Cantor's Paradox *}
   2.131  
   2.132  lemma Cantors_paradox [no_atp]:
     3.1 --- a/src/HOL/GCD.thy	Thu Aug 18 17:42:35 2011 +0200
     3.2 +++ b/src/HOL/GCD.thy	Thu Aug 18 22:50:28 2011 +0200
     3.3 @@ -531,11 +531,8 @@
     3.4  
     3.5  lemma Max_divisors_self_int[simp]: "n\<noteq>0 \<Longrightarrow> Max{d::int. d dvd n} = abs n"
     3.6  apply(rule antisym)
     3.7 - apply(rule Max_le_iff[THEN iffD2])
     3.8 -   apply simp
     3.9 -  apply fastsimp
    3.10 - apply (metis Collect_def abs_ge_self dvd_imp_le_int mem_def zle_trans)
    3.11 -apply simp
    3.12 + apply(rule Max_le_iff [THEN iffD2])
    3.13 +  apply (auto intro: abs_le_D1 dvd_imp_le_int)
    3.14  done
    3.15  
    3.16  lemma gcd_is_Max_divisors_nat:
     4.1 --- a/src/HOL/HOL.thy	Thu Aug 18 17:42:35 2011 +0200
     4.2 +++ b/src/HOL/HOL.thy	Thu Aug 18 22:50:28 2011 +0200
     4.3 @@ -1394,6 +1394,11 @@
     4.4    "f (if c then x else y) = (if c then f x else f y)"
     4.5    by simp
     4.6  
     4.7 +text{*As a simplification rule, it replaces all function equalities by
     4.8 +  first-order equalities.*}
     4.9 +lemma fun_eq_iff: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)"
    4.10 +  by auto
    4.11 +
    4.12  
    4.13  subsubsection {* Generic cases and induction *}
    4.14  
     5.1 --- a/src/HOL/IsaMakefile	Thu Aug 18 17:42:35 2011 +0200
     5.2 +++ b/src/HOL/IsaMakefile	Thu Aug 18 22:50:28 2011 +0200
     5.3 @@ -1039,32 +1039,30 @@
     5.4  
     5.5  $(LOG)/HOL-ex.gz: $(OUT)/HOL Decision_Procs/Commutative_Ring.thy	\
     5.6    Number_Theory/Primes.thy ex/Abstract_NAT.thy ex/Antiquote.thy		\
     5.7 -  ex/Arith_Examples.thy ex/Arithmetic_Series_Complex.thy \
     5.8 -  ex/BT.thy	ex/BinEx.thy ex/Binary.thy ex/Birthday_Paradox.thy			\
     5.9 -  ex/CTL.thy ex/Case_Product.thy			\
    5.10 -  ex/Chinese.thy ex/Classical.thy ex/CodegenSML_Test.thy		\
    5.11 -  ex/Coercion_Examples.thy ex/Coherent.thy ex/Dedekind_Real.thy		\
    5.12 -  ex/Efficient_Nat_examples.thy ex/Eval_Examples.thy ex/Fundefs.thy	\
    5.13 -  ex/Gauge_Integration.thy ex/Groebner_Examples.thy ex/Guess.thy	\
    5.14 -  ex/HarmonicSeries.thy ex/Hebrew.thy ex/Hex_Bin_Examples.thy		\
    5.15 -  ex/Higher_Order_Logic.thy ex/Iff_Oracle.thy ex/Induction_Schema.thy	\
    5.16 +  ex/Arith_Examples.thy ex/Arithmetic_Series_Complex.thy ex/BT.thy	\
    5.17 +  ex/BinEx.thy ex/Binary.thy ex/Birthday_Paradox.thy ex/CTL.thy		\
    5.18 +  ex/Case_Product.thy ex/Chinese.thy ex/Classical.thy			\
    5.19 +  ex/CodegenSML_Test.thy ex/Coercion_Examples.thy ex/Coherent.thy	\
    5.20 +  ex/Dedekind_Real.thy ex/Efficient_Nat_examples.thy			\
    5.21 +  ex/Eval_Examples.thy ex/Fundefs.thy ex/Gauge_Integration.thy		\
    5.22 +  ex/Groebner_Examples.thy ex/Guess.thy ex/HarmonicSeries.thy		\
    5.23 +  ex/Hebrew.thy ex/Hex_Bin_Examples.thy ex/Higher_Order_Logic.thy	\
    5.24 +  ex/Iff_Oracle.thy ex/Induction_Schema.thy				\
    5.25    ex/Interpretation_with_Defs.thy ex/Intuitionistic.thy ex/Lagrange.thy	\
    5.26    ex/List_to_Set_Comprehension_Examples.thy ex/LocaleTest2.thy		\
    5.27    ex/MT.thy ex/MergeSort.thy ex/Meson_Test.thy ex/MonoidGroup.thy	\
    5.28    ex/Multiquote.thy ex/NatSum.thy ex/Normalization_by_Evaluation.thy	\
    5.29    ex/Numeral.thy ex/PER.thy ex/PresburgerEx.thy ex/Primrec.thy		\
    5.30    ex/Quickcheck_Examples.thy ex/Quickcheck_Lattice_Examples.thy		\
    5.31 -  ex/Quickcheck_Narrowing_Examples.thy  				\
    5.32 -  ex/Quicksort.thy ex/ROOT.ML ex/Records.thy		\
    5.33 -  ex/ReflectionEx.thy ex/Refute_Examples.thy ex/SAT_Examples.thy	\
    5.34 -  ex/SVC_Oracle.thy ex/Serbian.thy ex/Set_Algebras.thy			\
    5.35 -  ex/sledgehammer_tactics.ML ex/Sqrt.thy				\
    5.36 -  ex/Sqrt_Script.thy ex/Sudoku.thy ex/Tarski.thy ex/Termination.thy	\
    5.37 -  ex/Transfer_Ex.thy ex/Tree23.thy			\
    5.38 +  ex/Quickcheck_Narrowing_Examples.thy ex/Quicksort.thy ex/ROOT.ML	\
    5.39 +  ex/Records.thy ex/ReflectionEx.thy ex/Refute_Examples.thy		\
    5.40 +  ex/SAT_Examples.thy ex/Serbian.thy ex/Set_Theory.thy			\
    5.41 +  ex/Set_Algebras.thy ex/SVC_Oracle.thy ex/sledgehammer_tactics.ML	\
    5.42 +  ex/Sqrt.thy ex/Sqrt_Script.thy ex/Sudoku.thy ex/Tarski.thy		\
    5.43 +  ex/Termination.thy ex/Transfer_Ex.thy ex/Tree23.thy			\
    5.44    ex/Unification.thy ex/While_Combinator_Example.thy			\
    5.45 -  ex/document/root.bib ex/document/root.tex ex/set.thy ex/svc_funcs.ML	\
    5.46 -  ex/svc_test.thy							\
    5.47 -  ../Tools/interpretation_with_defs.ML
    5.48 +  ex/document/root.bib ex/document/root.tex ex/svc_funcs.ML		\
    5.49 +  ex/svc_test.thy ../Tools/interpretation_with_defs.ML
    5.50  	@$(ISABELLE_TOOL) usedir $(OUT)/HOL ex
    5.51  
    5.52  
     6.1 --- a/src/HOL/Nat.thy	Thu Aug 18 17:42:35 2011 +0200
     6.2 +++ b/src/HOL/Nat.thy	Thu Aug 18 22:50:28 2011 +0200
     6.3 @@ -39,11 +39,20 @@
     6.4      Zero_RepI: "Nat Zero_Rep"
     6.5    | Suc_RepI: "Nat i \<Longrightarrow> Nat (Suc_Rep i)"
     6.6  
     6.7 -typedef (open Nat) nat = Nat
     6.8 -  by (rule exI, unfold mem_def, rule Nat.Zero_RepI)
     6.9 +typedef (open Nat) nat = "{n. Nat n}"
    6.10 +  using Nat.Zero_RepI by auto
    6.11  
    6.12 -definition Suc :: "nat => nat" where
    6.13 -  "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))"
    6.14 +lemma Nat_Rep_Nat:
    6.15 +  "Nat (Rep_Nat n)"
    6.16 +  using Rep_Nat by simp
    6.17 +
    6.18 +lemma Nat_Abs_Nat_inverse:
    6.19 +  "Nat n \<Longrightarrow> Rep_Nat (Abs_Nat n) = n"
    6.20 +  using Abs_Nat_inverse by simp
    6.21 +
    6.22 +lemma Nat_Abs_Nat_inject:
    6.23 +  "Nat n \<Longrightarrow> Nat m \<Longrightarrow> Abs_Nat n = Abs_Nat m \<longleftrightarrow> n = m"
    6.24 +  using Abs_Nat_inject by simp
    6.25  
    6.26  instantiation nat :: zero
    6.27  begin
    6.28 @@ -55,9 +64,11 @@
    6.29  
    6.30  end
    6.31  
    6.32 +definition Suc :: "nat \<Rightarrow> nat" where
    6.33 +  "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))"
    6.34 +
    6.35  lemma Suc_not_Zero: "Suc m \<noteq> 0"
    6.36 -  by (simp add: Zero_nat_def Suc_def Abs_Nat_inject [unfolded mem_def]
    6.37 -    Rep_Nat [unfolded mem_def] Suc_RepI Zero_RepI Suc_Rep_not_Zero_Rep [unfolded mem_def])
    6.38 +  by (simp add: Zero_nat_def Suc_def Suc_RepI Zero_RepI Nat_Abs_Nat_inject Suc_Rep_not_Zero_Rep Nat_Rep_Nat)
    6.39  
    6.40  lemma Zero_not_Suc: "0 \<noteq> Suc m"
    6.41    by (rule not_sym, rule Suc_not_Zero not_sym)
    6.42 @@ -67,12 +78,12 @@
    6.43  
    6.44  rep_datatype "0 \<Colon> nat" Suc
    6.45    apply (unfold Zero_nat_def Suc_def)
    6.46 -     apply (rule Rep_Nat_inverse [THEN subst]) -- {* types force good instantiation *}
    6.47 -     apply (erule Rep_Nat [unfolded mem_def, THEN Nat.induct])
    6.48 -     apply (iprover elim: Abs_Nat_inverse [unfolded mem_def, THEN subst])
    6.49 -    apply (simp_all add: Abs_Nat_inject [unfolded mem_def] Rep_Nat [unfolded mem_def]
    6.50 -      Suc_RepI Zero_RepI Suc_Rep_not_Zero_Rep [unfolded mem_def]
    6.51 -      Suc_Rep_not_Zero_Rep [unfolded mem_def, symmetric]
    6.52 +  apply (rule Rep_Nat_inverse [THEN subst]) -- {* types force good instantiation *}
    6.53 +   apply (erule Nat_Rep_Nat [THEN Nat.induct])
    6.54 +   apply (iprover elim: Nat_Abs_Nat_inverse [THEN subst])
    6.55 +    apply (simp_all add: Nat_Abs_Nat_inject Nat_Rep_Nat
    6.56 +      Suc_RepI Zero_RepI Suc_Rep_not_Zero_Rep
    6.57 +      Suc_Rep_not_Zero_Rep [symmetric]
    6.58        Suc_Rep_inject' Rep_Nat_inject)
    6.59    done
    6.60  
     7.1 --- a/src/HOL/Nitpick.thy	Thu Aug 18 17:42:35 2011 +0200
     7.2 +++ b/src/HOL/Nitpick.thy	Thu Aug 18 22:50:28 2011 +0200
     7.3 @@ -76,19 +76,19 @@
     7.4  definition prod :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set" where
     7.5  "prod A B = {(a, b). a \<in> A \<and> b \<in> B}"
     7.6  
     7.7 -definition refl' :: "('a \<times> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
     7.8 +definition refl' :: "('a \<times> 'a) set \<Rightarrow> bool" where
     7.9  "refl' r \<equiv> \<forall>x. (x, x) \<in> r"
    7.10  
    7.11 -definition wf' :: "('a \<times> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
    7.12 +definition wf' :: "('a \<times> 'a) set \<Rightarrow> bool" where
    7.13  "wf' r \<equiv> acyclic r \<and> (finite r \<or> unknown)"
    7.14  
    7.15 -definition card' :: "('a \<Rightarrow> bool) \<Rightarrow> nat" where
    7.16 +definition card' :: "'a set \<Rightarrow> nat" where
    7.17  "card' A \<equiv> if finite A then length (SOME xs. set xs = A \<and> distinct xs) else 0"
    7.18  
    7.19 -definition setsum' :: "('a \<Rightarrow> 'b\<Colon>comm_monoid_add) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'b" where
    7.20 +definition setsum' :: "('a \<Rightarrow> 'b\<Colon>comm_monoid_add) \<Rightarrow> 'a set \<Rightarrow> 'b" where
    7.21  "setsum' f A \<equiv> if finite A then listsum (map f (SOME xs. set xs = A \<and> distinct xs)) else 0"
    7.22  
    7.23 -inductive fold_graph' :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> bool" where
    7.24 +inductive fold_graph' :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool" where
    7.25  "fold_graph' f z {} z" |
    7.26  "\<lbrakk>x \<in> A; fold_graph' f z (A - {x}) y\<rbrakk> \<Longrightarrow> fold_graph' f z A (f x y)"
    7.27  
     8.1 --- a/src/HOL/RealDef.thy	Thu Aug 18 17:42:35 2011 +0200
     8.2 +++ b/src/HOL/RealDef.thy	Thu Aug 18 22:50:28 2011 +0200
     8.3 @@ -121,7 +121,7 @@
     8.4  subsection {* Cauchy sequences *}
     8.5  
     8.6  definition
     8.7 -  cauchy :: "(nat \<Rightarrow> rat) set"
     8.8 +  cauchy :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
     8.9  where
    8.10    "cauchy X \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r)"
    8.11  
     9.1 --- a/src/HOL/Relation.thy	Thu Aug 18 17:42:35 2011 +0200
     9.2 +++ b/src/HOL/Relation.thy	Thu Aug 18 22:50:28 2011 +0200
     9.3 @@ -133,9 +133,8 @@
     9.4  by blast
     9.5  
     9.6  lemma Id_on_def' [nitpick_unfold, code]:
     9.7 -  "(Id_on (A :: 'a => bool)) = (%(x, y). x = y \<and> A x)"
     9.8 -by (auto simp add: fun_eq_iff
     9.9 -  elim: Id_onE[unfolded mem_def] intro: Id_onI[unfolded mem_def])
    9.10 +  "Id_on {x. A x} = Collect (\<lambda>(x, y). x = y \<and> A x)"
    9.11 +by auto
    9.12  
    9.13  lemma Id_on_subset_Times: "Id_on A \<subseteq> A \<times> A"
    9.14  by blast
    10.1 --- a/src/HOL/String.thy	Thu Aug 18 17:42:35 2011 +0200
    10.2 +++ b/src/HOL/String.thy	Thu Aug 18 22:50:28 2011 +0200
    10.3 @@ -155,7 +155,7 @@
    10.4  
    10.5  subsection {* Strings as dedicated type *}
    10.6  
    10.7 -typedef (open) literal = "UNIV :: string \<Rightarrow> bool"
    10.8 +typedef (open) literal = "UNIV :: string set"
    10.9    morphisms explode STR ..
   10.10  
   10.11  instantiation literal :: size
    11.1 --- a/src/HOL/ex/ROOT.ML	Thu Aug 18 17:42:35 2011 +0200
    11.2 +++ b/src/HOL/ex/ROOT.ML	Thu Aug 18 22:50:28 2011 +0200
    11.3 @@ -48,7 +48,7 @@
    11.4    "Primrec",
    11.5    "Tarski",
    11.6    "Classical",
    11.7 -  "set",
    11.8 +  "Set_Theory",
    11.9    "Meson_Test",
   11.10    "Termination",
   11.11    "Coherent",
    12.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    12.2 +++ b/src/HOL/ex/Set_Theory.thy	Thu Aug 18 22:50:28 2011 +0200
    12.3 @@ -0,0 +1,227 @@
    12.4 +(*  Title:      HOL/ex/Set_Theory.thy
    12.5 +    Author:     Tobias Nipkow and Lawrence C Paulson
    12.6 +    Copyright   1991  University of Cambridge
    12.7 +*)
    12.8 +
    12.9 +header {* Set Theory examples: Cantor's Theorem, Schröder-Bernstein Theorem, etc. *}
   12.10 +
   12.11 +theory Set_Theory
   12.12 +imports Main
   12.13 +begin
   12.14 +
   12.15 +text{*
   12.16 +  These two are cited in Benzmueller and Kohlhase's system description
   12.17 +  of LEO, CADE-15, 1998 (pages 139-143) as theorems LEO could not
   12.18 +  prove.
   12.19 +*}
   12.20 +
   12.21 +lemma "(X = Y \<union> Z) =
   12.22 +    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
   12.23 +  by blast
   12.24 +
   12.25 +lemma "(X = Y \<inter> Z) =
   12.26 +    (X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))"
   12.27 +  by blast
   12.28 +
   12.29 +text {*
   12.30 +  Trivial example of term synthesis: apparently hard for some provers!
   12.31 +*}
   12.32 +
   12.33 +schematic_lemma "a \<noteq> b \<Longrightarrow> a \<in> ?X \<and> b \<notin> ?X"
   12.34 +  by blast
   12.35 +
   12.36 +
   12.37 +subsection {* Examples for the @{text blast} paper *}
   12.38 +
   12.39 +lemma "(\<Union>x \<in> C. f x \<union> g x) = \<Union>(f ` C)  \<union>  \<Union>(g ` C)"
   12.40 +  -- {* Union-image, called @{text Un_Union_image} in Main HOL *}
   12.41 +  by blast
   12.42 +
   12.43 +lemma "(\<Inter>x \<in> C. f x \<inter> g x) = \<Inter>(f ` C) \<inter> \<Inter>(g ` C)"
   12.44 +  -- {* Inter-image, called @{text Int_Inter_image} in Main HOL *}
   12.45 +  by blast
   12.46 +
   12.47 +lemma singleton_example_1:
   12.48 +     "\<And>S::'a set set. \<forall>x \<in> S. \<forall>y \<in> S. x \<subseteq> y \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
   12.49 +  by blast
   12.50 +
   12.51 +lemma singleton_example_2:
   12.52 +     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
   12.53 +  -- {*Variant of the problem above. *}
   12.54 +  by blast
   12.55 +
   12.56 +lemma "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"
   12.57 +  -- {* A unique fixpoint theorem --- @{text fast}/@{text best}/@{text meson} all fail. *}
   12.58 +  by metis
   12.59 +
   12.60 +
   12.61 +subsection {* Cantor's Theorem: There is no surjection from a set to its powerset *}
   12.62 +
   12.63 +lemma cantor1: "\<not> (\<exists>f:: 'a \<Rightarrow> 'a set. \<forall>S. \<exists>x. f x = S)"
   12.64 +  -- {* Requires best-first search because it is undirectional. *}
   12.65 +  by best
   12.66 +
   12.67 +schematic_lemma "\<forall>f:: 'a \<Rightarrow> 'a set. \<forall>x. f x \<noteq> ?S f"
   12.68 +  -- {*This form displays the diagonal term. *}
   12.69 +  by best
   12.70 +
   12.71 +schematic_lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
   12.72 +  -- {* This form exploits the set constructs. *}
   12.73 +  by (rule notI, erule rangeE, best)
   12.74 +
   12.75 +schematic_lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
   12.76 +  -- {* Or just this! *}
   12.77 +  by best
   12.78 +
   12.79 +
   12.80 +subsection {* The Schröder-Berstein Theorem *}
   12.81 +
   12.82 +lemma disj_lemma: "- (f ` X) = g ` (-X) \<Longrightarrow> f a = g b \<Longrightarrow> a \<in> X \<Longrightarrow> b \<in> X"
   12.83 +  by blast
   12.84 +
   12.85 +lemma surj_if_then_else:
   12.86 +  "-(f ` X) = g ` (-X) \<Longrightarrow> surj (\<lambda>z. if z \<in> X then f z else g z)"
   12.87 +  by (simp add: surj_def) blast
   12.88 +
   12.89 +lemma bij_if_then_else:
   12.90 +  "inj_on f X \<Longrightarrow> inj_on g (-X) \<Longrightarrow> -(f ` X) = g ` (-X) \<Longrightarrow>
   12.91 +    h = (\<lambda>z. if z \<in> X then f z else g z) \<Longrightarrow> inj h \<and> surj h"
   12.92 +  apply (unfold inj_on_def)
   12.93 +  apply (simp add: surj_if_then_else)
   12.94 +  apply (blast dest: disj_lemma sym)
   12.95 +  done
   12.96 +
   12.97 +lemma decomposition: "\<exists>X. X = - (g ` (- (f ` X)))"
   12.98 +  apply (rule exI)
   12.99 +  apply (rule lfp_unfold)
  12.100 +  apply (rule monoI, blast)
  12.101 +  done
  12.102 +
  12.103 +theorem Schroeder_Bernstein:
  12.104 +  "inj (f :: 'a \<Rightarrow> 'b) \<Longrightarrow> inj (g :: 'b \<Rightarrow> 'a)
  12.105 +    \<Longrightarrow> \<exists>h:: 'a \<Rightarrow> 'b. inj h \<and> surj h"
  12.106 +  apply (rule decomposition [where f=f and g=g, THEN exE])
  12.107 +  apply (rule_tac x = "(\<lambda>z. if z \<in> x then f z else inv g z)" in exI) 
  12.108 +    --{*The term above can be synthesized by a sufficiently detailed proof.*}
  12.109 +  apply (rule bij_if_then_else)
  12.110 +     apply (rule_tac [4] refl)
  12.111 +    apply (rule_tac [2] inj_on_inv_into)
  12.112 +    apply (erule subset_inj_on [OF _ subset_UNIV])
  12.113 +   apply blast
  12.114 +  apply (erule ssubst, subst double_complement, erule inv_image_comp [symmetric])
  12.115 +  done
  12.116 +
  12.117 +
  12.118 +subsection {* A simple party theorem *}
  12.119 +
  12.120 +text{* \emph{At any party there are two people who know the same
  12.121 +number of people}. Provided the party consists of at least two people
  12.122 +and the knows relation is symmetric. Knowing yourself does not count
  12.123 +--- otherwise knows needs to be reflexive. (From Freek Wiedijk's talk
  12.124 +at TPHOLs 2007.) *}
  12.125 +
  12.126 +lemma equal_number_of_acquaintances:
  12.127 +assumes "Domain R <= A" and "sym R" and "card A \<ge> 2"
  12.128 +shows "\<not> inj_on (%a. card(R `` {a} - {a})) A"
  12.129 +proof -
  12.130 +  let ?N = "%a. card(R `` {a} - {a})"
  12.131 +  let ?n = "card A"
  12.132 +  have "finite A" using `card A \<ge> 2` by(auto intro:ccontr)
  12.133 +  have 0: "R `` A <= A" using `sym R` `Domain R <= A`
  12.134 +    unfolding Domain_def sym_def by blast
  12.135 +  have h: "ALL a:A. R `` {a} <= A" using 0 by blast
  12.136 +  hence 1: "ALL a:A. finite(R `` {a})" using `finite A`
  12.137 +    by(blast intro: finite_subset)
  12.138 +  have sub: "?N ` A <= {0..<?n}"
  12.139 +  proof -
  12.140 +    have "ALL a:A. R `` {a} - {a} < A" using h by blast
  12.141 +    thus ?thesis using psubset_card_mono[OF `finite A`] by auto
  12.142 +  qed
  12.143 +  show "~ inj_on ?N A" (is "~ ?I")
  12.144 +  proof
  12.145 +    assume ?I
  12.146 +    hence "?n = card(?N ` A)" by(rule card_image[symmetric])
  12.147 +    with sub `finite A` have 2[simp]: "?N ` A = {0..<?n}"
  12.148 +      using subset_card_intvl_is_intvl[of _ 0] by(auto)
  12.149 +    have "0 : ?N ` A" and "?n - 1 : ?N ` A"  using `card A \<ge> 2` by simp+
  12.150 +    then obtain a b where ab: "a:A" "b:A" and Na: "?N a = 0" and Nb: "?N b = ?n - 1"
  12.151 +      by (auto simp del: 2)
  12.152 +    have "a \<noteq> b" using Na Nb `card A \<ge> 2` by auto
  12.153 +    have "R `` {a} - {a} = {}" by (metis 1 Na ab card_eq_0_iff finite_Diff)
  12.154 +    hence "b \<notin> R `` {a}" using `a\<noteq>b` by blast
  12.155 +    hence "a \<notin> R `` {b}" by (metis Image_singleton_iff assms(2) sym_def)
  12.156 +    hence 3: "R `` {b} - {b} <= A - {a,b}" using 0 ab by blast
  12.157 +    have 4: "finite (A - {a,b})" using `finite A` by simp
  12.158 +    have "?N b <= ?n - 2" using ab `a\<noteq>b` `finite A` card_mono[OF 4 3] by simp
  12.159 +    then show False using Nb `card A \<ge>  2` by arith
  12.160 +  qed
  12.161 +qed
  12.162 +
  12.163 +text {*
  12.164 +  From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages
  12.165 +  293-314.
  12.166 +
  12.167 +  Isabelle can prove the easy examples without any special mechanisms,
  12.168 +  but it can't prove the hard ones.
  12.169 +*}
  12.170 +
  12.171 +lemma "\<exists>A. (\<forall>x \<in> A. x \<le> (0::int))"
  12.172 +  -- {* Example 1, page 295. *}
  12.173 +  by force
  12.174 +
  12.175 +lemma "D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B"
  12.176 +  -- {* Example 2. *}
  12.177 +  by force
  12.178 +
  12.179 +lemma "P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)"
  12.180 +  -- {* Example 3. *}
  12.181 +  by force
  12.182 +
  12.183 +lemma "a < b \<and> b < (c::int) \<Longrightarrow> \<exists>A. a \<notin> A \<and> b \<in> A \<and> c \<notin> A"
  12.184 +  -- {* Example 4. *}
  12.185 +  by force
  12.186 +
  12.187 +lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
  12.188 +  -- {*Example 5, page 298. *}
  12.189 +  by force
  12.190 +
  12.191 +lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
  12.192 +  -- {* Example 6. *}
  12.193 +  by force
  12.194 +
  12.195 +lemma "\<exists>A. a \<notin> A"
  12.196 +  -- {* Example 7. *}
  12.197 +  by force
  12.198 +
  12.199 +lemma "(\<forall>u v. u < (0::int) \<longrightarrow> u \<noteq> abs v)
  12.200 +    \<longrightarrow> (\<exists>A::int set. (\<forall>y. abs y \<notin> A) \<and> -2 \<in> A)"
  12.201 +  -- {* Example 8 now needs a small hint. *}
  12.202 +  by (simp add: abs_if, force)
  12.203 +    -- {* not @{text blast}, which can't simplify @{text "-2 < 0"} *}
  12.204 +
  12.205 +text {* Example 9 omitted (requires the reals). *}
  12.206 +
  12.207 +text {* The paper has no Example 10! *}
  12.208 +
  12.209 +lemma "(\<forall>A. 0 \<in> A \<and> (\<forall>x \<in> A. Suc x \<in> A) \<longrightarrow> n \<in> A) \<and>
  12.210 +  P 0 \<and> (\<forall>x. P x \<longrightarrow> P (Suc x)) \<longrightarrow> P n"
  12.211 +  -- {* Example 11: needs a hint. *}
  12.212 +by(metis nat.induct)
  12.213 +
  12.214 +lemma
  12.215 +  "(\<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)
  12.216 +    \<and> P n \<longrightarrow> P m"
  12.217 +  -- {* Example 12. *}
  12.218 +  by auto
  12.219 +
  12.220 +lemma
  12.221 +  "(\<forall>x. (\<exists>u. x = 2 * u) = (\<not> (\<exists>v. Suc x = 2 * v))) \<longrightarrow>
  12.222 +    (\<exists>A. \<forall>x. (x \<in> A) = (Suc x \<notin> A))"
  12.223 +  -- {* Example EO1: typo in article, and with the obvious fix it seems
  12.224 +      to require arithmetic reasoning. *}
  12.225 +  apply clarify
  12.226 +  apply (rule_tac x = "{x. \<exists>u. x = 2 * u}" in exI, auto)
  12.227 +   apply metis+
  12.228 +  done
  12.229 +
  12.230 +end
    13.1 --- a/src/HOL/ex/set.thy	Thu Aug 18 17:42:35 2011 +0200
    13.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    13.3 @@ -1,225 +0,0 @@
    13.4 -(*  Title:      HOL/ex/set.thy
    13.5 -    Author:     Tobias Nipkow and Lawrence C Paulson
    13.6 -    Copyright   1991  University of Cambridge
    13.7 -*)
    13.8 -
    13.9 -header {* Set Theory examples: Cantor's Theorem, Schröder-Bernstein Theorem, etc. *}
   13.10 -
   13.11 -theory set imports Main begin
   13.12 -
   13.13 -text{*
   13.14 -  These two are cited in Benzmueller and Kohlhase's system description
   13.15 -  of LEO, CADE-15, 1998 (pages 139-143) as theorems LEO could not
   13.16 -  prove.
   13.17 -*}
   13.18 -
   13.19 -lemma "(X = Y \<union> Z) =
   13.20 -    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
   13.21 -  by blast
   13.22 -
   13.23 -lemma "(X = Y \<inter> Z) =
   13.24 -    (X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))"
   13.25 -  by blast
   13.26 -
   13.27 -text {*
   13.28 -  Trivial example of term synthesis: apparently hard for some provers!
   13.29 -*}
   13.30 -
   13.31 -schematic_lemma "a \<noteq> b \<Longrightarrow> a \<in> ?X \<and> b \<notin> ?X"
   13.32 -  by blast
   13.33 -
   13.34 -
   13.35 -subsection {* Examples for the @{text blast} paper *}
   13.36 -
   13.37 -lemma "(\<Union>x \<in> C. f x \<union> g x) = \<Union>(f ` C)  \<union>  \<Union>(g ` C)"
   13.38 -  -- {* Union-image, called @{text Un_Union_image} in Main HOL *}
   13.39 -  by blast
   13.40 -
   13.41 -lemma "(\<Inter>x \<in> C. f x \<inter> g x) = \<Inter>(f ` C) \<inter> \<Inter>(g ` C)"
   13.42 -  -- {* Inter-image, called @{text Int_Inter_image} in Main HOL *}
   13.43 -  by blast
   13.44 -
   13.45 -lemma singleton_example_1:
   13.46 -     "\<And>S::'a set set. \<forall>x \<in> S. \<forall>y \<in> S. x \<subseteq> y \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
   13.47 -  by blast
   13.48 -
   13.49 -lemma singleton_example_2:
   13.50 -     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
   13.51 -  -- {*Variant of the problem above. *}
   13.52 -  by blast
   13.53 -
   13.54 -lemma "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"
   13.55 -  -- {* A unique fixpoint theorem --- @{text fast}/@{text best}/@{text meson} all fail. *}
   13.56 -  by metis
   13.57 -
   13.58 -
   13.59 -subsection {* Cantor's Theorem: There is no surjection from a set to its powerset *}
   13.60 -
   13.61 -lemma cantor1: "\<not> (\<exists>f:: 'a \<Rightarrow> 'a set. \<forall>S. \<exists>x. f x = S)"
   13.62 -  -- {* Requires best-first search because it is undirectional. *}
   13.63 -  by best
   13.64 -
   13.65 -schematic_lemma "\<forall>f:: 'a \<Rightarrow> 'a set. \<forall>x. f x \<noteq> ?S f"
   13.66 -  -- {*This form displays the diagonal term. *}
   13.67 -  by best
   13.68 -
   13.69 -schematic_lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
   13.70 -  -- {* This form exploits the set constructs. *}
   13.71 -  by (rule notI, erule rangeE, best)
   13.72 -
   13.73 -schematic_lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
   13.74 -  -- {* Or just this! *}
   13.75 -  by best
   13.76 -
   13.77 -
   13.78 -subsection {* The Schröder-Berstein Theorem *}
   13.79 -
   13.80 -lemma disj_lemma: "- (f ` X) = g ` (-X) \<Longrightarrow> f a = g b \<Longrightarrow> a \<in> X \<Longrightarrow> b \<in> X"
   13.81 -  by blast
   13.82 -
   13.83 -lemma surj_if_then_else:
   13.84 -  "-(f ` X) = g ` (-X) \<Longrightarrow> surj (\<lambda>z. if z \<in> X then f z else g z)"
   13.85 -  by (simp add: surj_def) blast
   13.86 -
   13.87 -lemma bij_if_then_else:
   13.88 -  "inj_on f X \<Longrightarrow> inj_on g (-X) \<Longrightarrow> -(f ` X) = g ` (-X) \<Longrightarrow>
   13.89 -    h = (\<lambda>z. if z \<in> X then f z else g z) \<Longrightarrow> inj h \<and> surj h"
   13.90 -  apply (unfold inj_on_def)
   13.91 -  apply (simp add: surj_if_then_else)
   13.92 -  apply (blast dest: disj_lemma sym)
   13.93 -  done
   13.94 -
   13.95 -lemma decomposition: "\<exists>X. X = - (g ` (- (f ` X)))"
   13.96 -  apply (rule exI)
   13.97 -  apply (rule lfp_unfold)
   13.98 -  apply (rule monoI, blast)
   13.99 -  done
  13.100 -
  13.101 -theorem Schroeder_Bernstein:
  13.102 -  "inj (f :: 'a \<Rightarrow> 'b) \<Longrightarrow> inj (g :: 'b \<Rightarrow> 'a)
  13.103 -    \<Longrightarrow> \<exists>h:: 'a \<Rightarrow> 'b. inj h \<and> surj h"
  13.104 -  apply (rule decomposition [where f=f and g=g, THEN exE])
  13.105 -  apply (rule_tac x = "(\<lambda>z. if z \<in> x then f z else inv g z)" in exI) 
  13.106 -    --{*The term above can be synthesized by a sufficiently detailed proof.*}
  13.107 -  apply (rule bij_if_then_else)
  13.108 -     apply (rule_tac [4] refl)
  13.109 -    apply (rule_tac [2] inj_on_inv_into)
  13.110 -    apply (erule subset_inj_on [OF _ subset_UNIV])
  13.111 -   apply blast
  13.112 -  apply (erule ssubst, subst double_complement, erule inv_image_comp [symmetric])
  13.113 -  done
  13.114 -
  13.115 -
  13.116 -subsection {* A simple party theorem *}
  13.117 -
  13.118 -text{* \emph{At any party there are two people who know the same
  13.119 -number of people}. Provided the party consists of at least two people
  13.120 -and the knows relation is symmetric. Knowing yourself does not count
  13.121 ---- otherwise knows needs to be reflexive. (From Freek Wiedijk's talk
  13.122 -at TPHOLs 2007.) *}
  13.123 -
  13.124 -lemma equal_number_of_acquaintances:
  13.125 -assumes "Domain R <= A" and "sym R" and "card A \<ge> 2"
  13.126 -shows "\<not> inj_on (%a. card(R `` {a} - {a})) A"
  13.127 -proof -
  13.128 -  let ?N = "%a. card(R `` {a} - {a})"
  13.129 -  let ?n = "card A"
  13.130 -  have "finite A" using `card A \<ge> 2` by(auto intro:ccontr)
  13.131 -  have 0: "R `` A <= A" using `sym R` `Domain R <= A`
  13.132 -    unfolding Domain_def sym_def by blast
  13.133 -  have h: "ALL a:A. R `` {a} <= A" using 0 by blast
  13.134 -  hence 1: "ALL a:A. finite(R `` {a})" using `finite A`
  13.135 -    by(blast intro: finite_subset)
  13.136 -  have sub: "?N ` A <= {0..<?n}"
  13.137 -  proof -
  13.138 -    have "ALL a:A. R `` {a} - {a} < A" using h by blast
  13.139 -    thus ?thesis using psubset_card_mono[OF `finite A`] by auto
  13.140 -  qed
  13.141 -  show "~ inj_on ?N A" (is "~ ?I")
  13.142 -  proof
  13.143 -    assume ?I
  13.144 -    hence "?n = card(?N ` A)" by(rule card_image[symmetric])
  13.145 -    with sub `finite A` have 2[simp]: "?N ` A = {0..<?n}"
  13.146 -      using subset_card_intvl_is_intvl[of _ 0] by(auto)
  13.147 -    have "0 : ?N ` A" and "?n - 1 : ?N ` A"  using `card A \<ge> 2` by simp+
  13.148 -    then obtain a b where ab: "a:A" "b:A" and Na: "?N a = 0" and Nb: "?N b = ?n - 1"
  13.149 -      by (auto simp del: 2)
  13.150 -    have "a \<noteq> b" using Na Nb `card A \<ge> 2` by auto
  13.151 -    have "R `` {a} - {a} = {}" by (metis 1 Na ab card_eq_0_iff finite_Diff)
  13.152 -    hence "b \<notin> R `` {a}" using `a\<noteq>b` by blast
  13.153 -    hence "a \<notin> R `` {b}" by (metis Image_singleton_iff assms(2) sym_def)
  13.154 -    hence 3: "R `` {b} - {b} <= A - {a,b}" using 0 ab by blast
  13.155 -    have 4: "finite (A - {a,b})" using `finite A` by simp
  13.156 -    have "?N b <= ?n - 2" using ab `a\<noteq>b` `finite A` card_mono[OF 4 3] by simp
  13.157 -    then show False using Nb `card A \<ge>  2` by arith
  13.158 -  qed
  13.159 -qed
  13.160 -
  13.161 -text {*
  13.162 -  From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages
  13.163 -  293-314.
  13.164 -
  13.165 -  Isabelle can prove the easy examples without any special mechanisms,
  13.166 -  but it can't prove the hard ones.
  13.167 -*}
  13.168 -
  13.169 -lemma "\<exists>A. (\<forall>x \<in> A. x \<le> (0::int))"
  13.170 -  -- {* Example 1, page 295. *}
  13.171 -  by force
  13.172 -
  13.173 -lemma "D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B"
  13.174 -  -- {* Example 2. *}
  13.175 -  by force
  13.176 -
  13.177 -lemma "P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)"
  13.178 -  -- {* Example 3. *}
  13.179 -  by force
  13.180 -
  13.181 -lemma "a < b \<and> b < (c::int) \<Longrightarrow> \<exists>A. a \<notin> A \<and> b \<in> A \<and> c \<notin> A"
  13.182 -  -- {* Example 4. *}
  13.183 -  by force
  13.184 -
  13.185 -lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
  13.186 -  -- {*Example 5, page 298. *}
  13.187 -  by force
  13.188 -
  13.189 -lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
  13.190 -  -- {* Example 6. *}
  13.191 -  by force
  13.192 -
  13.193 -lemma "\<exists>A. a \<notin> A"
  13.194 -  -- {* Example 7. *}
  13.195 -  by force
  13.196 -
  13.197 -lemma "(\<forall>u v. u < (0::int) \<longrightarrow> u \<noteq> abs v)
  13.198 -    \<longrightarrow> (\<exists>A::int set. (\<forall>y. abs y \<notin> A) \<and> -2 \<in> A)"
  13.199 -  -- {* Example 8 now needs a small hint. *}
  13.200 -  by (simp add: abs_if, force)
  13.201 -    -- {* not @{text blast}, which can't simplify @{text "-2 < 0"} *}
  13.202 -
  13.203 -text {* Example 9 omitted (requires the reals). *}
  13.204 -
  13.205 -text {* The paper has no Example 10! *}
  13.206 -
  13.207 -lemma "(\<forall>A. 0 \<in> A \<and> (\<forall>x \<in> A. Suc x \<in> A) \<longrightarrow> n \<in> A) \<and>
  13.208 -  P 0 \<and> (\<forall>x. P x \<longrightarrow> P (Suc x)) \<longrightarrow> P n"
  13.209 -  -- {* Example 11: needs a hint. *}
  13.210 -by(metis nat.induct)
  13.211 -
  13.212 -lemma
  13.213 -  "(\<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)
  13.214 -    \<and> P n \<longrightarrow> P m"
  13.215 -  -- {* Example 12. *}
  13.216 -  by auto
  13.217 -
  13.218 -lemma
  13.219 -  "(\<forall>x. (\<exists>u. x = 2 * u) = (\<not> (\<exists>v. Suc x = 2 * v))) \<longrightarrow>
  13.220 -    (\<exists>A. \<forall>x. (x \<in> A) = (Suc x \<notin> A))"
  13.221 -  -- {* Example EO1: typo in article, and with the obvious fix it seems
  13.222 -      to require arithmetic reasoning. *}
  13.223 -  apply clarify
  13.224 -  apply (rule_tac x = "{x. \<exists>u. x = 2 * u}" in exI, auto)
  13.225 -   apply metis+
  13.226 -  done
  13.227 -
  13.228 -end