| author | wenzelm | 
| Thu, 18 Jul 2024 11:02:08 +0200 | |
| changeset 80582 | 1bc7eef961ff | 
| parent 80399 | 413a86331bf6 | 
| permissions | -rw-r--r-- | 
| 44276 | 1  | 
(* Title: HOL/ex/Set_Theory.thy  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
2  | 
Author: Tobias Nipkow and Lawrence C Paulson  | 
| 
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
3  | 
Copyright 1991 University of Cambridge  | 
| 13107 | 4  | 
*)  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
5  | 
|
| 61343 | 6  | 
section \<open>Set Theory examples: Cantor's Theorem, Schröder-Bernstein Theorem, etc.\<close>  | 
| 9100 | 7  | 
|
| 44276 | 8  | 
theory Set_Theory  | 
9  | 
imports Main  | 
|
10  | 
begin  | 
|
| 9100 | 11  | 
|
| 61343 | 12  | 
text\<open>  | 
| 13107 | 13  | 
These two are cited in Benzmueller and Kohlhase's system description  | 
14  | 
of LEO, CADE-15, 1998 (pages 139-143) as theorems LEO could not  | 
|
15  | 
prove.  | 
|
| 61343 | 16  | 
\<close>  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
17  | 
|
| 13107 | 18  | 
lemma "(X = Y \<union> Z) =  | 
19  | 
(Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"  | 
|
20  | 
by blast  | 
|
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
21  | 
|
| 13107 | 22  | 
lemma "(X = Y \<inter> Z) =  | 
23  | 
(X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))"  | 
|
24  | 
by blast  | 
|
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
25  | 
|
| 61343 | 26  | 
text \<open>  | 
| 13107 | 27  | 
Trivial example of term synthesis: apparently hard for some provers!  | 
| 61343 | 28  | 
\<close>  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
29  | 
|
| 61337 | 30  | 
schematic_goal "a \<noteq> b \<Longrightarrow> a \<in> ?X \<and> b \<notin> ?X"  | 
| 13107 | 31  | 
by blast  | 
32  | 
||
33  | 
||
| 61933 | 34  | 
subsection \<open>Examples for the \<open>blast\<close> paper\<close>  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
35  | 
|
| 13107 | 36  | 
lemma "(\<Union>x \<in> C. f x \<union> g x) = \<Union>(f ` C) \<union> \<Union>(g ` C)"  | 
| 61933 | 37  | 
\<comment> \<open>Union-image, called \<open>Un_Union_image\<close> in Main HOL\<close>  | 
| 13107 | 38  | 
by blast  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
39  | 
|
| 13107 | 40  | 
lemma "(\<Inter>x \<in> C. f x \<inter> g x) = \<Inter>(f ` C) \<inter> \<Inter>(g ` C)"  | 
| 61933 | 41  | 
\<comment> \<open>Inter-image, called \<open>Int_Inter_image\<close> in Main HOL\<close>  | 
| 13107 | 42  | 
by blast  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
43  | 
|
| 16898 | 44  | 
lemma singleton_example_1:  | 
45  | 
     "\<And>S::'a set set. \<forall>x \<in> S. \<forall>y \<in> S. x \<subseteq> y \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
 | 
|
| 18391 | 46  | 
by blast  | 
| 16898 | 47  | 
|
48  | 
lemma singleton_example_2:  | 
|
49  | 
     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
 | 
|
| 61933 | 50  | 
\<comment> \<open>Variant of the problem above.\<close>  | 
| 18391 | 51  | 
by blast  | 
| 13107 | 52  | 
|
53  | 
lemma "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"  | 
|
| 61933 | 54  | 
\<comment> \<open>A unique fixpoint theorem --- \<open>fast\<close>/\<open>best\<close>/\<open>meson\<close> all fail.\<close>  | 
| 24573 | 55  | 
by metis  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
56  | 
|
| 
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
57  | 
|
| 61343 | 58  | 
subsection \<open>Cantor's Theorem: There is no surjection from a set to its powerset\<close>  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
59  | 
|
| 13107 | 60  | 
lemma cantor1: "\<not> (\<exists>f:: 'a \<Rightarrow> 'a set. \<forall>S. \<exists>x. f x = S)"  | 
| 61933 | 61  | 
\<comment> \<open>Requires best-first search because it is undirectional.\<close>  | 
| 13107 | 62  | 
by best  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
63  | 
|
| 61337 | 64  | 
schematic_goal "\<forall>f:: 'a \<Rightarrow> 'a set. \<forall>x. f x \<noteq> ?S f"  | 
| 61933 | 65  | 
\<comment> \<open>This form displays the diagonal term.\<close>  | 
| 13107 | 66  | 
by best  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
67  | 
|
| 61337 | 68  | 
schematic_goal "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"  | 
| 61933 | 69  | 
\<comment> \<open>This form exploits the set constructs.\<close>  | 
| 13107 | 70  | 
by (rule notI, erule rangeE, best)  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
71  | 
|
| 61337 | 72  | 
schematic_goal "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"  | 
| 61933 | 73  | 
\<comment> \<open>Or just this!\<close>  | 
| 13107 | 74  | 
by best  | 
75  | 
||
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
76  | 
|
| 61937 | 77  | 
subsection \<open>The Schröder-Bernstein Theorem\<close>  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
78  | 
|
| 80399 | 79  | 
lemma decomposition:  | 
80  | 
obtains X where "X = - (g ` (- (f ` X)))"  | 
|
81  | 
using lfp_unfold [OF monoI, of "\<lambda>X. - g ` (- f ` X)"]  | 
|
| 13107 | 82  | 
by blast  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
83  | 
|
| 13107 | 84  | 
theorem Schroeder_Bernstein:  | 
| 80399 | 85  | 
fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'a"  | 
86  | 
assumes "inj f" "inj g"  | 
|
87  | 
obtains h:: "'a \<Rightarrow> 'b" where "inj h" "surj h"  | 
|
88  | 
proof (rule decomposition)  | 
|
89  | 
fix X  | 
|
90  | 
assume X: "X = - (g ` (- (f ` X)))"  | 
|
91  | 
let ?h = "\<lambda>z. if z \<in> X then f z else inv g z"  | 
|
92  | 
show thesis  | 
|
93  | 
proof  | 
|
94  | 
have "inj_on (inv g) (-X)"  | 
|
95  | 
by (metis X \<open>inj g\<close> bij_betw_def double_complement inj_imp_bij_betw_inv)  | 
|
96  | 
with \<open>inj f\<close> show "inj ?h"  | 
|
97  | 
unfolding inj_on_def by (metis Compl_iff X \<open>inj g\<close> imageI image_inv_f_f)  | 
|
98  | 
show "surj ?h"  | 
|
99  | 
using \<open>inj g\<close> X image_iff surj_def by fastforce  | 
|
100  | 
qed  | 
|
101  | 
qed  | 
|
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
102  | 
|
| 61343 | 103  | 
subsection \<open>A simple party theorem\<close>  | 
| 24853 | 104  | 
|
| 61343 | 105  | 
text\<open>\emph{At any party there are two people who know the same
 | 
| 24853 | 106  | 
number of people}. Provided the party consists of at least two people  | 
107  | 
and the knows relation is symmetric. Knowing yourself does not count  | 
|
108  | 
--- otherwise knows needs to be reflexive. (From Freek Wiedijk's talk  | 
|
| 61343 | 109  | 
at TPHOLs 2007.)\<close>  | 
| 24853 | 110  | 
|
111  | 
lemma equal_number_of_acquaintances:  | 
|
112  | 
assumes "Domain R <= A" and "sym R" and "card A \<ge> 2"  | 
|
113  | 
shows "\<not> inj_on (%a. card(R `` {a} - {a})) A"
 | 
|
114  | 
proof -  | 
|
115  | 
  let ?N = "%a. card(R `` {a} - {a})"
 | 
|
116  | 
let ?n = "card A"  | 
|
| 61343 | 117  | 
have "finite A" using \<open>card A \<ge> 2\<close> by(auto intro:ccontr)  | 
118  | 
have 0: "R `` A <= A" using \<open>sym R\<close> \<open>Domain R <= A\<close>  | 
|
| 
46752
 
e9e7209eb375
more fundamental pred-to-set conversions, particularly by means of inductive_set; associated consolidation of some theorem names (c.f. NEWS)
 
haftmann 
parents: 
45966 
diff
changeset
 | 
119  | 
unfolding Domain_unfold sym_def by blast  | 
| 67613 | 120  | 
  have h: "\<forall>a\<in>A. R `` {a} <= A" using 0 by blast
 | 
121  | 
  hence 1: "\<forall>a\<in>A. finite(R `` {a})" using \<open>finite A\<close>
 | 
|
| 24853 | 122  | 
by(blast intro: finite_subset)  | 
123  | 
  have sub: "?N ` A <= {0..<?n}"
 | 
|
124  | 
proof -  | 
|
| 67613 | 125  | 
    have "\<forall>a\<in>A. R `` {a} - {a} < A" using h by blast
 | 
| 61343 | 126  | 
thus ?thesis using psubset_card_mono[OF \<open>finite A\<close>] by auto  | 
| 24853 | 127  | 
qed  | 
128  | 
show "~ inj_on ?N A" (is "~ ?I")  | 
|
129  | 
proof  | 
|
130  | 
assume ?I  | 
|
131  | 
hence "?n = card(?N ` A)" by(rule card_image[symmetric])  | 
|
| 61343 | 132  | 
    with sub \<open>finite A\<close> have 2[simp]: "?N ` A = {0..<?n}"
 | 
| 24853 | 133  | 
using subset_card_intvl_is_intvl[of _ 0] by(auto)  | 
| 67613 | 134  | 
have "0 \<in> ?N ` A" and "?n - 1 \<in> ?N ` A" using \<open>card A \<ge> 2\<close> by simp+  | 
135  | 
then obtain a b where ab: "a\<in>A" "b\<in>A" and Na: "?N a = 0" and Nb: "?N b = ?n - 1"  | 
|
| 24853 | 136  | 
by (auto simp del: 2)  | 
| 61343 | 137  | 
have "a \<noteq> b" using Na Nb \<open>card A \<ge> 2\<close> by auto  | 
| 24853 | 138  | 
    have "R `` {a} - {a} = {}" by (metis 1 Na ab card_eq_0_iff finite_Diff)
 | 
| 61343 | 139  | 
    hence "b \<notin> R `` {a}" using \<open>a\<noteq>b\<close> by blast
 | 
| 24853 | 140  | 
    hence "a \<notin> R `` {b}" by (metis Image_singleton_iff assms(2) sym_def)
 | 
141  | 
    hence 3: "R `` {b} - {b} <= A - {a,b}" using 0 ab by blast
 | 
|
| 61343 | 142  | 
    have 4: "finite (A - {a,b})" using \<open>finite A\<close> by simp
 | 
143  | 
have "?N b <= ?n - 2" using ab \<open>a\<noteq>b\<close> \<open>finite A\<close> card_mono[OF 4 3] by simp  | 
|
144  | 
then show False using Nb \<open>card A \<ge> 2\<close> by arith  | 
|
| 24853 | 145  | 
qed  | 
146  | 
qed  | 
|
147  | 
||
| 61343 | 148  | 
text \<open>  | 
| 13107 | 149  | 
From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages  | 
150  | 
293-314.  | 
|
151  | 
||
152  | 
Isabelle can prove the easy examples without any special mechanisms,  | 
|
153  | 
but it can't prove the hard ones.  | 
|
| 61343 | 154  | 
\<close>  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
155  | 
|
| 13107 | 156  | 
lemma "\<exists>A. (\<forall>x \<in> A. x \<le> (0::int))"  | 
| 61933 | 157  | 
\<comment> \<open>Example 1, page 295.\<close>  | 
| 13107 | 158  | 
by force  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
159  | 
|
| 13107 | 160  | 
lemma "D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B"  | 
| 61933 | 161  | 
\<comment> \<open>Example 2.\<close>  | 
| 13107 | 162  | 
by force  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
163  | 
|
| 13107 | 164  | 
lemma "P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)"  | 
| 61933 | 165  | 
\<comment> \<open>Example 3.\<close>  | 
| 13107 | 166  | 
by force  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
167  | 
|
| 13107 | 168  | 
lemma "a < b \<and> b < (c::int) \<Longrightarrow> \<exists>A. a \<notin> A \<and> b \<in> A \<and> c \<notin> A"  | 
| 61933 | 169  | 
\<comment> \<open>Example 4.\<close>  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
63804 
diff
changeset
 | 
170  | 
by auto \<comment> \<open>slow\<close>  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
171  | 
|
| 13107 | 172  | 
lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"  | 
| 61933 | 173  | 
\<comment> \<open>Example 5, page 298.\<close>  | 
| 13107 | 174  | 
by force  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
175  | 
|
| 13107 | 176  | 
lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"  | 
| 61933 | 177  | 
\<comment> \<open>Example 6.\<close>  | 
| 13107 | 178  | 
by force  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
179  | 
|
| 13107 | 180  | 
lemma "\<exists>A. a \<notin> A"  | 
| 61933 | 181  | 
\<comment> \<open>Example 7.\<close>  | 
| 13107 | 182  | 
by force  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
183  | 
|
| 61945 | 184  | 
lemma "(\<forall>u v. u < (0::int) \<longrightarrow> u \<noteq> \<bar>v\<bar>)  | 
185  | 
\<longrightarrow> (\<exists>A::int set. -2 \<in> A & (\<forall>y. \<bar>y\<bar> \<notin> A))"  | 
|
| 61933 | 186  | 
\<comment> \<open>Example 8 needs a small hint.\<close>  | 
| 45966 | 187  | 
by force  | 
| 61933 | 188  | 
\<comment> \<open>not \<open>blast\<close>, which can't simplify \<open>-2 < 0\<close>\<close>  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
189  | 
|
| 61343 | 190  | 
text \<open>Example 9 omitted (requires the reals).\<close>  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
191  | 
|
| 61343 | 192  | 
text \<open>The paper has no Example 10!\<close>  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
193  | 
|
| 13107 | 194  | 
lemma "(\<forall>A. 0 \<in> A \<and> (\<forall>x \<in> A. Suc x \<in> A) \<longrightarrow> n \<in> A) \<and>  | 
195  | 
P 0 \<and> (\<forall>x. P x \<longrightarrow> P (Suc x)) \<longrightarrow> P n"  | 
|
| 61933 | 196  | 
\<comment> \<open>Example 11: needs a hint.\<close>  | 
| 
40928
 
ace26e2cee91
eliminated unqualified accesses of datatype facts -- it seems like they all of them were unintended
 
krauss 
parents: 
36319 
diff
changeset
 | 
197  | 
by(metis nat.induct)  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
198  | 
|
| 13107 | 199  | 
lemma  | 
200  | 
"(\<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)  | 
|
201  | 
\<and> P n \<longrightarrow> P m"  | 
|
| 61933 | 202  | 
\<comment> \<open>Example 12.\<close>  | 
| 13107 | 203  | 
by auto  | 
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
204  | 
|
| 13107 | 205  | 
lemma  | 
206  | 
"(\<forall>x. (\<exists>u. x = 2 * u) = (\<not> (\<exists>v. Suc x = 2 * v))) \<longrightarrow>  | 
|
207  | 
(\<exists>A. \<forall>x. (x \<in> A) = (Suc x \<notin> A))"  | 
|
| 61933 | 208  | 
\<comment> \<open>Example EO1: typo in article, and with the obvious fix it seems  | 
| 80399 | 209  | 
to require arithmetic reasoning. 2024-06-19: now trivial for sledgehammer (LCP)\<close>  | 
210  | 
by (metis even_Suc mem_Collect_eq)  | 
|
| 
13058
 
ad6106d7b4bb
converted theory "set" to Isar and added some SET-VAR examples
 
paulson 
parents: 
9100 
diff
changeset
 | 
211  | 
|
| 9100 | 212  | 
end  |