author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Tue, 20 Apr 2010 14:56:20 +0200 | |
changeset 36215 | 88ff48884d26 |
parent 34055 | fdf294ee08b2 |
child 36319 | 8feb2c4bef1a |
permissions | -rw-r--r-- |
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
1 |
(* Title: HOL/ex/set.thy |
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
2 |
ID: $Id$ |
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
3 |
Author: Tobias Nipkow and Lawrence C Paulson |
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
4 |
Copyright 1991 University of Cambridge |
13107 | 5 |
*) |
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
6 |
|
19982 | 7 |
header {* Set Theory examples: Cantor's Theorem, Schröder-Bernstein Theorem, etc. *} |
9100 | 8 |
|
16417 | 9 |
theory set imports Main begin |
9100 | 10 |
|
13107 | 11 |
text{* |
12 |
These two are cited in Benzmueller and Kohlhase's system description |
|
13 |
of LEO, CADE-15, 1998 (pages 139-143) as theorems LEO could not |
|
14 |
prove. |
|
15 |
*} |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
16 |
|
13107 | 17 |
lemma "(X = Y \<union> Z) = |
18 |
(Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" |
|
19 |
by blast |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
20 |
|
13107 | 21 |
lemma "(X = Y \<inter> Z) = |
22 |
(X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))" |
|
23 |
by blast |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
24 |
|
13107 | 25 |
text {* |
26 |
Trivial example of term synthesis: apparently hard for some provers! |
|
27 |
*} |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
28 |
|
13107 | 29 |
lemma "a \<noteq> b \<Longrightarrow> a \<in> ?X \<and> b \<notin> ?X" |
30 |
by blast |
|
31 |
||
32 |
||
33 |
subsection {* Examples for the @{text blast} paper *} |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
34 |
|
13107 | 35 |
lemma "(\<Union>x \<in> C. f x \<union> g x) = \<Union>(f ` C) \<union> \<Union>(g ` C)" |
36 |
-- {* Union-image, called @{text Un_Union_image} in Main HOL *} |
|
37 |
by blast |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
38 |
|
13107 | 39 |
lemma "(\<Inter>x \<in> C. f x \<inter> g x) = \<Inter>(f ` C) \<inter> \<Inter>(g ` C)" |
40 |
-- {* Inter-image, called @{text Int_Inter_image} in Main HOL *} |
|
41 |
by blast |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
42 |
|
16898 | 43 |
lemma singleton_example_1: |
44 |
"\<And>S::'a set set. \<forall>x \<in> S. \<forall>y \<in> S. x \<subseteq> y \<Longrightarrow> \<exists>z. S \<subseteq> {z}" |
|
18391 | 45 |
by blast |
16898 | 46 |
|
47 |
lemma singleton_example_2: |
|
48 |
"\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}" |
|
49 |
-- {*Variant of the problem above. *} |
|
18391 | 50 |
by blast |
13107 | 51 |
|
52 |
lemma "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y" |
|
53 |
-- {* A unique fixpoint theorem --- @{text fast}/@{text best}/@{text meson} all fail. *} |
|
24573 | 54 |
by metis |
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
55 |
|
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
56 |
|
13107 | 57 |
subsection {* Cantor's Theorem: There is no surjection from a set to its powerset *} |
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
58 |
|
13107 | 59 |
lemma cantor1: "\<not> (\<exists>f:: 'a \<Rightarrow> 'a set. \<forall>S. \<exists>x. f x = S)" |
60 |
-- {* Requires best-first search because it is undirectional. *} |
|
61 |
by best |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
62 |
|
13107 | 63 |
lemma "\<forall>f:: 'a \<Rightarrow> 'a set. \<forall>x. f x \<noteq> ?S f" |
64 |
-- {*This form displays the diagonal term. *} |
|
65 |
by best |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
66 |
|
13107 | 67 |
lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)" |
68 |
-- {* This form exploits the set constructs. *} |
|
69 |
by (rule notI, erule rangeE, best) |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
70 |
|
13107 | 71 |
lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)" |
72 |
-- {* Or just this! *} |
|
73 |
by best |
|
74 |
||
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
75 |
|
13107 | 76 |
subsection {* The Schröder-Berstein Theorem *} |
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
77 |
|
13107 | 78 |
lemma disj_lemma: "- (f ` X) = g ` (-X) \<Longrightarrow> f a = g b \<Longrightarrow> a \<in> X \<Longrightarrow> b \<in> X" |
79 |
by blast |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
80 |
|
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
81 |
lemma surj_if_then_else: |
13107 | 82 |
"-(f ` X) = g ` (-X) \<Longrightarrow> surj (\<lambda>z. if z \<in> X then f z else g z)" |
83 |
by (simp add: surj_def) blast |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
84 |
|
13107 | 85 |
lemma bij_if_then_else: |
86 |
"inj_on f X \<Longrightarrow> inj_on g (-X) \<Longrightarrow> -(f ` X) = g ` (-X) \<Longrightarrow> |
|
87 |
h = (\<lambda>z. if z \<in> X then f z else g z) \<Longrightarrow> inj h \<and> surj h" |
|
88 |
apply (unfold inj_on_def) |
|
89 |
apply (simp add: surj_if_then_else) |
|
90 |
apply (blast dest: disj_lemma sym) |
|
91 |
done |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
92 |
|
13107 | 93 |
lemma decomposition: "\<exists>X. X = - (g ` (- (f ` X)))" |
94 |
apply (rule exI) |
|
95 |
apply (rule lfp_unfold) |
|
96 |
apply (rule monoI, blast) |
|
97 |
done |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
98 |
|
13107 | 99 |
theorem Schroeder_Bernstein: |
100 |
"inj (f :: 'a \<Rightarrow> 'b) \<Longrightarrow> inj (g :: 'b \<Rightarrow> 'a) |
|
101 |
\<Longrightarrow> \<exists>h:: 'a \<Rightarrow> 'b. inj h \<and> surj h" |
|
15488 | 102 |
apply (rule decomposition [where f=f and g=g, THEN exE]) |
103 |
apply (rule_tac x = "(\<lambda>z. if z \<in> x then f z else inv g z)" in exI) |
|
104 |
--{*The term above can be synthesized by a sufficiently detailed proof.*} |
|
13107 | 105 |
apply (rule bij_if_then_else) |
106 |
apply (rule_tac [4] refl) |
|
33057 | 107 |
apply (rule_tac [2] inj_on_inv_into) |
15306 | 108 |
apply (erule subset_inj_on [OF _ subset_UNIV]) |
15488 | 109 |
apply blast |
110 |
apply (erule ssubst, subst double_complement, erule inv_image_comp [symmetric]) |
|
13107 | 111 |
done |
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
112 |
|
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
113 |
|
24853 | 114 |
subsection {* A simple party theorem *} |
115 |
||
116 |
text{* \emph{At any party there are two people who know the same |
|
117 |
number of people}. Provided the party consists of at least two people |
|
118 |
and the knows relation is symmetric. Knowing yourself does not count |
|
119 |
--- otherwise knows needs to be reflexive. (From Freek Wiedijk's talk |
|
120 |
at TPHOLs 2007.) *} |
|
121 |
||
122 |
lemma equal_number_of_acquaintances: |
|
123 |
assumes "Domain R <= A" and "sym R" and "card A \<ge> 2" |
|
124 |
shows "\<not> inj_on (%a. card(R `` {a} - {a})) A" |
|
125 |
proof - |
|
126 |
let ?N = "%a. card(R `` {a} - {a})" |
|
127 |
let ?n = "card A" |
|
128 |
have "finite A" using `card A \<ge> 2` by(auto intro:ccontr) |
|
129 |
have 0: "R `` A <= A" using `sym R` `Domain R <= A` |
|
130 |
unfolding Domain_def sym_def by blast |
|
131 |
have h: "ALL a:A. R `` {a} <= A" using 0 by blast |
|
132 |
hence 1: "ALL a:A. finite(R `` {a})" using `finite A` |
|
133 |
by(blast intro: finite_subset) |
|
134 |
have sub: "?N ` A <= {0..<?n}" |
|
135 |
proof - |
|
136 |
have "ALL a:A. R `` {a} - {a} < A" using h by blast |
|
137 |
thus ?thesis using psubset_card_mono[OF `finite A`] by auto |
|
138 |
qed |
|
139 |
show "~ inj_on ?N A" (is "~ ?I") |
|
140 |
proof |
|
141 |
assume ?I |
|
142 |
hence "?n = card(?N ` A)" by(rule card_image[symmetric]) |
|
143 |
with sub `finite A` have 2[simp]: "?N ` A = {0..<?n}" |
|
144 |
using subset_card_intvl_is_intvl[of _ 0] by(auto) |
|
145 |
have "0 : ?N ` A" and "?n - 1 : ?N ` A" using `card A \<ge> 2` by simp+ |
|
146 |
then obtain a b where ab: "a:A" "b:A" and Na: "?N a = 0" and Nb: "?N b = ?n - 1" |
|
147 |
by (auto simp del: 2) |
|
148 |
have "a \<noteq> b" using Na Nb `card A \<ge> 2` by auto |
|
149 |
have "R `` {a} - {a} = {}" by (metis 1 Na ab card_eq_0_iff finite_Diff) |
|
150 |
hence "b \<notin> R `` {a}" using `a\<noteq>b` by blast |
|
151 |
hence "a \<notin> R `` {b}" by (metis Image_singleton_iff assms(2) sym_def) |
|
152 |
hence 3: "R `` {b} - {b} <= A - {a,b}" using 0 ab by blast |
|
153 |
have 4: "finite (A - {a,b})" using `finite A` by simp |
|
154 |
have "?N b <= ?n - 2" using ab `a\<noteq>b` `finite A` card_mono[OF 4 3] by simp |
|
155 |
then show False using Nb `card A \<ge> 2` by arith |
|
156 |
qed |
|
157 |
qed |
|
158 |
||
13107 | 159 |
text {* |
160 |
From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages |
|
161 |
293-314. |
|
162 |
||
163 |
Isabelle can prove the easy examples without any special mechanisms, |
|
164 |
but it can't prove the hard ones. |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
165 |
*} |
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
166 |
|
13107 | 167 |
lemma "\<exists>A. (\<forall>x \<in> A. x \<le> (0::int))" |
168 |
-- {* Example 1, page 295. *} |
|
169 |
by force |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
170 |
|
13107 | 171 |
lemma "D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B" |
172 |
-- {* Example 2. *} |
|
173 |
by force |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
174 |
|
13107 | 175 |
lemma "P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)" |
176 |
-- {* Example 3. *} |
|
177 |
by force |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
178 |
|
13107 | 179 |
lemma "a < b \<and> b < (c::int) \<Longrightarrow> \<exists>A. a \<notin> A \<and> b \<in> A \<and> c \<notin> A" |
180 |
-- {* Example 4. *} |
|
181 |
by force |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
182 |
|
13107 | 183 |
lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A" |
184 |
-- {*Example 5, page 298. *} |
|
185 |
by force |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
186 |
|
13107 | 187 |
lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A" |
188 |
-- {* Example 6. *} |
|
189 |
by force |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
190 |
|
13107 | 191 |
lemma "\<exists>A. a \<notin> A" |
192 |
-- {* Example 7. *} |
|
193 |
by force |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
194 |
|
13107 | 195 |
lemma "(\<forall>u v. u < (0::int) \<longrightarrow> u \<noteq> abs v) |
196 |
\<longrightarrow> (\<exists>A::int set. (\<forall>y. abs y \<notin> A) \<and> -2 \<in> A)" |
|
14353
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
13107
diff
changeset
|
197 |
-- {* Example 8 now needs a small hint. *} |
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
13107
diff
changeset
|
198 |
by (simp add: abs_if, force) |
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
13107
diff
changeset
|
199 |
-- {* not @{text blast}, which can't simplify @{text "-2 < 0"} *} |
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
200 |
|
13107 | 201 |
text {* Example 9 omitted (requires the reals). *} |
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
202 |
|
13107 | 203 |
text {* The paper has no Example 10! *} |
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
204 |
|
13107 | 205 |
lemma "(\<forall>A. 0 \<in> A \<and> (\<forall>x \<in> A. Suc x \<in> A) \<longrightarrow> n \<in> A) \<and> |
206 |
P 0 \<and> (\<forall>x. P x \<longrightarrow> P (Suc x)) \<longrightarrow> P n" |
|
207 |
-- {* Example 11: needs a hint. *} |
|
34055 | 208 |
by(metis Nat.induct) |
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
209 |
|
13107 | 210 |
lemma |
211 |
"(\<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) |
|
212 |
\<and> P n \<longrightarrow> P m" |
|
213 |
-- {* Example 12. *} |
|
214 |
by auto |
|
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
215 |
|
13107 | 216 |
lemma |
217 |
"(\<forall>x. (\<exists>u. x = 2 * u) = (\<not> (\<exists>v. Suc x = 2 * v))) \<longrightarrow> |
|
218 |
(\<exists>A. \<forall>x. (x \<in> A) = (Suc x \<notin> A))" |
|
219 |
-- {* Example EO1: typo in article, and with the obvious fix it seems |
|
220 |
to require arithmetic reasoning. *} |
|
221 |
apply clarify |
|
222 |
apply (rule_tac x = "{x. \<exists>u. x = 2 * u}" in exI, auto) |
|
34055 | 223 |
apply metis+ |
13107 | 224 |
done |
13058
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
paulson
parents:
9100
diff
changeset
|
225 |
|
9100 | 226 |
end |