(* Title: HOL/ex/set.thy
ID: $Id: set.thy,v 1.9 2005/07/20 15:01:20 paulson Exp $
Author: Tobias Nipkow and Lawrence C Paulson
Copyright 1991 University of Cambridge
*)
header {* Set Theory examples: Cantor's Theorem, Schröder-Berstein 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!
*}
lemma "a ≠ b ==> a ∈ ?X ∧ b ∉ ?X"
by blast
subsection {* Examples for the @{text blast} paper *}
lemma "(\<Union>x ∈ C. f x ∪ g x) = \<Union>(f ` C) ∪ \<Union>(g ` C)"
-- {* Union-image, called @{text Un_Union_image} in Main HOL *}
by blast
lemma "(\<Inter>x ∈ C. f x ∩ g x) = \<Inter>(f ` C) ∩ \<Inter>(g ` C)"
-- {* Inter-image, called @{text Int_Inter_image} in Main HOL *}
by blast
text{*Both of the singleton examples can be proved very quickly by @{text
"blast del: UNIV_I"} but not by @{text blast} alone. For some reason, @{text
UNIV_I} greatly increases the search space.*}
lemma singleton_example_1:
"!!S::'a set set. ∀x ∈ S. ∀y ∈ S. x ⊆ y ==> ∃z. S ⊆ {z}"
by (meson subsetI subset_antisym insertCI)
lemma singleton_example_2:
"∀x ∈ S. \<Union>S ⊆ x ==> ∃z. S ⊆ {z}"
-- {*Variant of the problem above. *}
by (meson subsetI subset_antisym insertCI UnionI)
lemma "∃!x. f (g x) = x ==> ∃!y. g (f y) = y"
-- {* A unique fixpoint theorem --- @{text fast}/@{text best}/@{text meson} all fail. *}
apply (erule ex1E, rule ex1I, erule arg_cong)
apply (rule subst, assumption, erule allE, rule arg_cong, erule mp)
apply (erule arg_cong)
done
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
lemma "∀f:: 'a => 'a set. ∀x. f x ≠ ?S f"
-- {*This form displays the diagonal term. *}
by best
lemma "?S ∉ range (f :: 'a => 'a set)"
-- {* This form exploits the set constructs. *}
by (rule notI, erule rangeE, best)
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)
apply (erule subset_inj_on [OF _ subset_UNIV])
apply blast
apply (erule ssubst, subst double_complement, erule inv_image_comp [symmetric])
done
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. *}
apply clarify
apply (drule_tac x = "{x. P x}" in spec)
apply force
done
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 (case_tac v, auto)
apply (drule_tac x = "Suc v" and P = "λx. ?a x ≠ ?b x" in spec, force)
done
end
lemma
(X = Y ∪ Z) = (Y ⊆ X ∧ Z ⊆ X ∧ (∀V. Y ⊆ V ∧ Z ⊆ V --> X ⊆ V))
lemma
(X = Y ∩ Z) = (X ⊆ Y ∧ X ⊆ Z ∧ (∀V. V ⊆ Y ∧ V ⊆ Z --> V ⊆ X))
lemma
a ≠ b ==> a ∈ {a} ∧ b ∉ {a}
lemma
(UN x:C. f x ∪ g x) = Union (f ` C) ∪ Union (g ` C)
lemma
(INT x:C. f x ∩ g x) = Inter (f ` C) ∩ Inter (g ` C)
lemma singleton_example_1:
∀x∈S. ∀y∈S. x ⊆ y ==> ∃z. S ⊆ {z}
lemma singleton_example_2:
∀x∈S. Union S ⊆ x ==> ∃z. S ⊆ {z}
lemma
∃!x. f (g x) = x ==> ∃!y. g (f y) = y
lemma cantor1:
¬ (∃f. ∀S. ∃x. f x = S)
lemma
∀f x. f x ≠ {x. x ∉ f x}
lemma
{x. x ∉ f x} ∉ range f
lemma
{x. x ∉ f x} ∉ range f
lemma disj_lemma:
[| - f ` X = g ` (- X); f a = g b; a ∈ X |] ==> b ∈ X
lemma surj_if_then_else:
- f ` X = g ` (- X) ==> surj (%z. if z ∈ X then f z else g z)
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
lemma decomposition:
∃X. X = - g ` (- f ` X)
theorem Schroeder_Bernstein:
[| inj f; inj g |] ==> ∃h. inj h ∧ surj h
lemma
∃A. ∀x∈A. x ≤ 0
lemma
D ∈ F ==> ∃G. ∀A∈G. ∃B∈F. A ⊆ B
lemma
P a ==> ∃A. (∀x∈A. P x) ∧ (∃y. y ∈ A)
lemma
a < b ∧ b < c ==> ∃A. a ∉ A ∧ b ∈ A ∧ c ∉ A
lemma
P (f b) ==> ∃s A. (∀x∈A. P x) ∧ f s ∈ A
lemma
P (f b) ==> ∃s A. (∀x∈A. P x) ∧ f s ∈ A
lemma
∃A. a ∉ A
lemma
(∀u v. u < 0 --> u ≠ ¦v¦) --> (∃A. (∀y. ¦y¦ ∉ A) ∧ -2 ∈ A)
lemma
(∀A. 0 ∈ A ∧ (∀x∈A. Suc x ∈ A) --> n ∈ A) ∧ P 0 ∧ (∀x. P x --> P (Suc x)) --> P n
lemma
(∀A. (0, 0) ∈ A ∧ (∀x y. (x, y) ∈ A --> (Suc x, Suc y) ∈ A) --> (n, m) ∈ A) ∧ P n --> P m
lemma
(∀x. (∃u. x = 2 * u) = (¬ (∃v. Suc x = 2 * v))) --> (∃A. ∀x. (x ∈ A) = (Suc x ∉ A))