# HG changeset patch # User haftmann # Date 1313700628 -7200 # Node ID d27b9fe4759e8bc08a35e1d94086d86d5a52f4a8 # Parent d995733b635df12acc85729801ac546a9f52b45f# Parent 4846f3f320d9fbd90c354b474e2e1a74268367d3 merged diff -r d995733b635d -r d27b9fe4759e src/HOL/Equiv_Relations.thy --- a/src/HOL/Equiv_Relations.thy Thu Aug 18 17:42:35 2011 +0200 +++ b/src/HOL/Equiv_Relations.thy Thu Aug 18 22:50:28 2011 +0200 @@ -164,7 +164,7 @@ text{*A congruence-preserving function*} -definition congruent :: "('a \ 'a \ bool) \ ('a \ 'b) \ bool" where +definition congruent :: "('a \ 'a) set \ ('a \ 'b) \ bool" where "congruent r f \ (\(y, z) \ r. f y = f z)" lemma congruentI: @@ -229,7 +229,7 @@ text{*A congruence-preserving function of two arguments*} -definition congruent2 :: "('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b \ 'c) \ bool" where +definition congruent2 :: "('a \ 'a) set \ ('b \ 'b) set \ ('a \ 'b \ 'c) \ bool" where "congruent2 r1 r2 f \ (\(y1, z1) \ r1. \(y2, z2) \ r2. f y1 y2 = f z1 z2)" lemma congruent2I': diff -r d995733b635d -r d27b9fe4759e src/HOL/Fun.thy --- a/src/HOL/Fun.thy Thu Aug 18 17:42:35 2011 +0200 +++ b/src/HOL/Fun.thy Thu Aug 18 22:50:28 2011 +0200 @@ -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 \ (\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 \ (\x. P x \ g (f x) = x) \ P x \ x = g u" by auto @@ -26,26 +17,22 @@ subsection {* The Identity Function @{text id} *} -definition - id :: "'a \ 'a" -where +definition id :: "'a \ 'a" where "id = (\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 \ g"} *} -definition - comp :: "('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" (infixl "o" 55) -where +definition comp :: "('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" (infixl "o" 55) where "f o g = (\x. f (g x))" notation (xsymbols) @@ -54,9 +41,6 @@ notation (HTML output) comp (infixl "\" 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 \ 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 \ ((\v. a (b v) = c (d v)) \ R) \ R" @@ -89,9 +73,7 @@ subsection {* The Forward Composition Operator @{text fcomp} *} -definition - fcomp :: "('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" (infixl "\>" 60) -where +definition fcomp :: "('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" (infixl "\>" 60) where "f \> g = (\x. g (f x))" lemma fcomp_apply [simp]: "(f \> 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 \ 'b) \ ('a \ 'b) \ 'a set \ 'a \ 'b" -where +definition override_on :: "('a \ 'b) \ ('a \ 'b) \ 'a set \ 'a \ 'b" where "override_on f g A = (\a. if a \ 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 \ 'a \ ('a \ 'b) \ ('a \ 'b)" -where +definition swap :: "'a \ 'a \ ('a \ 'b) \ ('a \ '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]: diff -r d995733b635d -r d27b9fe4759e src/HOL/GCD.thy --- a/src/HOL/GCD.thy Thu Aug 18 17:42:35 2011 +0200 +++ b/src/HOL/GCD.thy Thu Aug 18 22:50:28 2011 +0200 @@ -531,11 +531,8 @@ lemma Max_divisors_self_int[simp]: "n\0 \ 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: diff -r d995733b635d -r d27b9fe4759e src/HOL/HOL.thy --- a/src/HOL/HOL.thy Thu Aug 18 17:42:35 2011 +0200 +++ b/src/HOL/HOL.thy Thu Aug 18 22:50:28 2011 +0200 @@ -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 \ (\x. f x = g x)" + by auto + subsubsection {* Generic cases and induction *} diff -r d995733b635d -r d27b9fe4759e src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Thu Aug 18 17:42:35 2011 +0200 +++ b/src/HOL/IsaMakefile Thu Aug 18 22:50:28 2011 +0200 @@ -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 diff -r d995733b635d -r d27b9fe4759e src/HOL/Nat.thy --- a/src/HOL/Nat.thy Thu Aug 18 17:42:35 2011 +0200 +++ b/src/HOL/Nat.thy Thu Aug 18 22:50:28 2011 +0200 @@ -39,11 +39,20 @@ Zero_RepI: "Nat Zero_Rep" | Suc_RepI: "Nat i \ 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 -definition Suc :: "nat => nat" where - "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))" +lemma Nat_Rep_Nat: + "Nat (Rep_Nat n)" + using Rep_Nat by simp + +lemma Nat_Abs_Nat_inverse: + "Nat n \ Rep_Nat (Abs_Nat n) = n" + using Abs_Nat_inverse by simp + +lemma Nat_Abs_Nat_inject: + "Nat n \ Nat m \ Abs_Nat n = Abs_Nat m \ n = m" + using Abs_Nat_inject by simp instantiation nat :: zero begin @@ -55,9 +64,11 @@ end +definition Suc :: "nat \ nat" where + "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))" + lemma Suc_not_Zero: "Suc m \ 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 \ Suc m" by (rule not_sym, rule Suc_not_Zero not_sym) @@ -67,12 +78,12 @@ rep_datatype "0 \ 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 diff -r d995733b635d -r d27b9fe4759e src/HOL/Nitpick.thy --- a/src/HOL/Nitpick.thy Thu Aug 18 17:42:35 2011 +0200 +++ b/src/HOL/Nitpick.thy Thu Aug 18 22:50:28 2011 +0200 @@ -76,19 +76,19 @@ definition prod :: "'a set \ 'b set \ ('a \ 'b) set" where "prod A B = {(a, b). a \ A \ b \ B}" -definition refl' :: "('a \ 'a \ bool) \ bool" where +definition refl' :: "('a \ 'a) set \ bool" where "refl' r \ \x. (x, x) \ r" -definition wf' :: "('a \ 'a \ bool) \ bool" where +definition wf' :: "('a \ 'a) set \ bool" where "wf' r \ acyclic r \ (finite r \ unknown)" -definition card' :: "('a \ bool) \ nat" where +definition card' :: "'a set \ nat" where "card' A \ if finite A then length (SOME xs. set xs = A \ distinct xs) else 0" -definition setsum' :: "('a \ 'b\comm_monoid_add) \ ('a \ bool) \ 'b" where +definition setsum' :: "('a \ 'b\comm_monoid_add) \ 'a set \ 'b" where "setsum' f A \ if finite A then listsum (map f (SOME xs. set xs = A \ distinct xs)) else 0" -inductive fold_graph' :: "('a \ 'b \ 'b) \ 'b \ ('a \ bool) \ 'b \ bool" where +inductive fold_graph' :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b \ bool" where "fold_graph' f z {} z" | "\x \ A; fold_graph' f z (A - {x}) y\ \ fold_graph' f z A (f x y)" diff -r d995733b635d -r d27b9fe4759e src/HOL/RealDef.thy --- a/src/HOL/RealDef.thy Thu Aug 18 17:42:35 2011 +0200 +++ b/src/HOL/RealDef.thy Thu Aug 18 22:50:28 2011 +0200 @@ -121,7 +121,7 @@ subsection {* Cauchy sequences *} definition - cauchy :: "(nat \ rat) set" + cauchy :: "(nat \ rat) \ bool" where "cauchy X \ (\r>0. \k. \m\k. \n\k. \X m - X n\ < r)" diff -r d995733b635d -r d27b9fe4759e src/HOL/Relation.thy --- a/src/HOL/Relation.thy Thu Aug 18 17:42:35 2011 +0200 +++ b/src/HOL/Relation.thy Thu Aug 18 22:50:28 2011 +0200 @@ -133,9 +133,8 @@ by blast lemma Id_on_def' [nitpick_unfold, code]: - "(Id_on (A :: 'a => bool)) = (%(x, y). x = y \ 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 (\(x, y). x = y \ A x)" +by auto lemma Id_on_subset_Times: "Id_on A \ A \ A" by blast diff -r d995733b635d -r d27b9fe4759e src/HOL/String.thy --- a/src/HOL/String.thy Thu Aug 18 17:42:35 2011 +0200 +++ b/src/HOL/String.thy Thu Aug 18 22:50:28 2011 +0200 @@ -155,7 +155,7 @@ subsection {* Strings as dedicated type *} -typedef (open) literal = "UNIV :: string \ bool" +typedef (open) literal = "UNIV :: string set" morphisms explode STR .. instantiation literal :: size diff -r d995733b635d -r d27b9fe4759e src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Thu Aug 18 17:42:35 2011 +0200 +++ b/src/HOL/ex/ROOT.ML Thu Aug 18 22:50:28 2011 +0200 @@ -48,7 +48,7 @@ "Primrec", "Tarski", "Classical", - "set", + "Set_Theory", "Meson_Test", "Termination", "Coherent", diff -r d995733b635d -r d27b9fe4759e src/HOL/ex/Set_Theory.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Set_Theory.thy Thu Aug 18 22:50:28 2011 +0200 @@ -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 \ Z) = + (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" + by blast + +lemma "(X = Y \ Z) = + (X \ Y \ X \ Z \ (\V. V \ Y \ V \ Z \ V \ X))" + by blast + +text {* + Trivial example of term synthesis: apparently hard for some provers! +*} + +schematic_lemma "a \ b \ a \ ?X \ b \ ?X" + by blast + + +subsection {* Examples for the @{text blast} paper *} + +lemma "(\x \ C. f x \ g x) = \(f ` C) \ \(g ` C)" + -- {* Union-image, called @{text Un_Union_image} in Main HOL *} + by blast + +lemma "(\x \ C. f x \ g x) = \(f ` C) \ \(g ` C)" + -- {* Inter-image, called @{text Int_Inter_image} in Main HOL *} + by blast + +lemma singleton_example_1: + "\S::'a set set. \x \ S. \y \ S. x \ y \ \z. S \ {z}" + by blast + +lemma singleton_example_2: + "\x \ S. \S \ x \ \z. S \ {z}" + -- {*Variant of the problem above. *} + by blast + +lemma "\!x. f (g x) = x \ \!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: "\ (\f:: 'a \ 'a set. \S. \x. f x = S)" + -- {* Requires best-first search because it is undirectional. *} + by best + +schematic_lemma "\f:: 'a \ 'a set. \x. f x \ ?S f" + -- {*This form displays the diagonal term. *} + by best + +schematic_lemma "?S \ range (f :: 'a \ 'a set)" + -- {* This form exploits the set constructs. *} + by (rule notI, erule rangeE, best) + +schematic_lemma "?S \ range (f :: 'a \ 'a set)" + -- {* Or just this! *} + by best + + +subsection {* The Schröder-Berstein Theorem *} + +lemma disj_lemma: "- (f ` X) = g ` (-X) \ f a = g b \ a \ X \ b \ X" + by blast + +lemma surj_if_then_else: + "-(f ` X) = g ` (-X) \ surj (\z. if z \ X then f z else g z)" + by (simp add: surj_def) blast + +lemma bij_if_then_else: + "inj_on f X \ inj_on g (-X) \ -(f ` X) = g ` (-X) \ + h = (\z. if z \ X then f z else g z) \ inj h \ surj h" + apply (unfold inj_on_def) + apply (simp add: surj_if_then_else) + apply (blast dest: disj_lemma sym) + done + +lemma decomposition: "\X. X = - (g ` (- (f ` X)))" + apply (rule exI) + apply (rule lfp_unfold) + apply (rule monoI, blast) + done + +theorem Schroeder_Bernstein: + "inj (f :: 'a \ 'b) \ inj (g :: 'b \ 'a) + \ \h:: 'a \ 'b. inj h \ surj h" + apply (rule decomposition [where f=f and g=g, THEN exE]) + apply (rule_tac x = "(\z. if z \ 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 \ 2" +shows "\ 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 \ 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.. 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 \ b" using Na Nb `card A \ 2` by auto + have "R `` {a} - {a} = {}" by (metis 1 Na ab card_eq_0_iff finite_Diff) + hence "b \ R `` {a}" using `a\b` by blast + hence "a \ 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\b` `finite A` card_mono[OF 4 3] by simp + then show False using Nb `card A \ 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 "\A. (\x \ A. x \ (0::int))" + -- {* Example 1, page 295. *} + by force + +lemma "D \ F \ \G. \A \ G. \B \ F. A \ B" + -- {* Example 2. *} + by force + +lemma "P a \ \A. (\x \ A. P x) \ (\y. y \ A)" + -- {* Example 3. *} + by force + +lemma "a < b \ b < (c::int) \ \A. a \ A \ b \ A \ c \ A" + -- {* Example 4. *} + by force + +lemma "P (f b) \ \s A. (\x \ A. P x) \ f s \ A" + -- {*Example 5, page 298. *} + by force + +lemma "P (f b) \ \s A. (\x \ A. P x) \ f s \ A" + -- {* Example 6. *} + by force + +lemma "\A. a \ A" + -- {* Example 7. *} + by force + +lemma "(\u v. u < (0::int) \ u \ abs v) + \ (\A::int set. (\y. abs y \ A) \ -2 \ 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 "(\A. 0 \ A \ (\x \ A. Suc x \ A) \ n \ A) \ + P 0 \ (\x. P x \ P (Suc x)) \ P n" + -- {* Example 11: needs a hint. *} +by(metis nat.induct) + +lemma + "(\A. (0, 0) \ A \ (\x y. (x, y) \ A \ (Suc x, Suc y) \ A) \ (n, m) \ A) + \ P n \ P m" + -- {* Example 12. *} + by auto + +lemma + "(\x. (\u. x = 2 * u) = (\ (\v. Suc x = 2 * v))) \ + (\A. \x. (x \ A) = (Suc x \ A))" + -- {* Example EO1: typo in article, and with the obvious fix it seems + to require arithmetic reasoning. *} + apply clarify + apply (rule_tac x = "{x. \u. x = 2 * u}" in exI, auto) + apply metis+ + done + +end diff -r d995733b635d -r d27b9fe4759e src/HOL/ex/set.thy --- a/src/HOL/ex/set.thy Thu Aug 18 17:42:35 2011 +0200 +++ /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 \ Z) = - (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" - by blast - -lemma "(X = Y \ Z) = - (X \ Y \ X \ Z \ (\V. V \ Y \ V \ Z \ V \ X))" - by blast - -text {* - Trivial example of term synthesis: apparently hard for some provers! -*} - -schematic_lemma "a \ b \ a \ ?X \ b \ ?X" - by blast - - -subsection {* Examples for the @{text blast} paper *} - -lemma "(\x \ C. f x \ g x) = \(f ` C) \ \(g ` C)" - -- {* Union-image, called @{text Un_Union_image} in Main HOL *} - by blast - -lemma "(\x \ C. f x \ g x) = \(f ` C) \ \(g ` C)" - -- {* Inter-image, called @{text Int_Inter_image} in Main HOL *} - by blast - -lemma singleton_example_1: - "\S::'a set set. \x \ S. \y \ S. x \ y \ \z. S \ {z}" - by blast - -lemma singleton_example_2: - "\x \ S. \S \ x \ \z. S \ {z}" - -- {*Variant of the problem above. *} - by blast - -lemma "\!x. f (g x) = x \ \!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: "\ (\f:: 'a \ 'a set. \S. \x. f x = S)" - -- {* Requires best-first search because it is undirectional. *} - by best - -schematic_lemma "\f:: 'a \ 'a set. \x. f x \ ?S f" - -- {*This form displays the diagonal term. *} - by best - -schematic_lemma "?S \ range (f :: 'a \ 'a set)" - -- {* This form exploits the set constructs. *} - by (rule notI, erule rangeE, best) - -schematic_lemma "?S \ range (f :: 'a \ 'a set)" - -- {* Or just this! *} - by best - - -subsection {* The Schröder-Berstein Theorem *} - -lemma disj_lemma: "- (f ` X) = g ` (-X) \ f a = g b \ a \ X \ b \ X" - by blast - -lemma surj_if_then_else: - "-(f ` X) = g ` (-X) \ surj (\z. if z \ X then f z else g z)" - by (simp add: surj_def) blast - -lemma bij_if_then_else: - "inj_on f X \ inj_on g (-X) \ -(f ` X) = g ` (-X) \ - h = (\z. if z \ X then f z else g z) \ inj h \ surj h" - apply (unfold inj_on_def) - apply (simp add: surj_if_then_else) - apply (blast dest: disj_lemma sym) - done - -lemma decomposition: "\X. X = - (g ` (- (f ` X)))" - apply (rule exI) - apply (rule lfp_unfold) - apply (rule monoI, blast) - done - -theorem Schroeder_Bernstein: - "inj (f :: 'a \ 'b) \ inj (g :: 'b \ 'a) - \ \h:: 'a \ 'b. inj h \ surj h" - apply (rule decomposition [where f=f and g=g, THEN exE]) - apply (rule_tac x = "(\z. if z \ 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 \ 2" -shows "\ 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 \ 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.. 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 \ b" using Na Nb `card A \ 2` by auto - have "R `` {a} - {a} = {}" by (metis 1 Na ab card_eq_0_iff finite_Diff) - hence "b \ R `` {a}" using `a\b` by blast - hence "a \ 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\b` `finite A` card_mono[OF 4 3] by simp - then show False using Nb `card A \ 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 "\A. (\x \ A. x \ (0::int))" - -- {* Example 1, page 295. *} - by force - -lemma "D \ F \ \G. \A \ G. \B \ F. A \ B" - -- {* Example 2. *} - by force - -lemma "P a \ \A. (\x \ A. P x) \ (\y. y \ A)" - -- {* Example 3. *} - by force - -lemma "a < b \ b < (c::int) \ \A. a \ A \ b \ A \ c \ A" - -- {* Example 4. *} - by force - -lemma "P (f b) \ \s A. (\x \ A. P x) \ f s \ A" - -- {*Example 5, page 298. *} - by force - -lemma "P (f b) \ \s A. (\x \ A. P x) \ f s \ A" - -- {* Example 6. *} - by force - -lemma "\A. a \ A" - -- {* Example 7. *} - by force - -lemma "(\u v. u < (0::int) \ u \ abs v) - \ (\A::int set. (\y. abs y \ A) \ -2 \ 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 "(\A. 0 \ A \ (\x \ A. Suc x \ A) \ n \ A) \ - P 0 \ (\x. P x \ P (Suc x)) \ P n" - -- {* Example 11: needs a hint. *} -by(metis nat.induct) - -lemma - "(\A. (0, 0) \ A \ (\x y. (x, y) \ A \ (Suc x, Suc y) \ A) \ (n, m) \ A) - \ P n \ P m" - -- {* Example 12. *} - by auto - -lemma - "(\x. (\u. x = 2 * u) = (\ (\v. Suc x = 2 * v))) \ - (\A. \x. (x \ A) = (Suc x \ A))" - -- {* Example EO1: typo in article, and with the obvious fix it seems - to require arithmetic reasoning. *} - apply clarify - apply (rule_tac x = "{x. \u. x = 2 * u}" in exI, auto) - apply metis+ - done - -end